src/HOL/Nominal/nominal_package.ML
author berghofe
Wed, 28 Mar 2007 19:17:40 +0200
changeset 22543 9460a4cf0acc
parent 22529 902ed60d53a7
child 22578 b0eb5652f210
permissions -rw-r--r--
Renamed induct(s)_weak to weak_induct(s) in order to unify naming scheme with the one used in nominal_inductive.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_package.ML
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     2
    ID:         $Id$
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     3
    Author:     Stefan Berghofer and Christian Urban, TU Muenchen
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     4
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     5
Nominal datatype package for Isabelle/HOL.
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     6
*)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     7
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     8
signature NOMINAL_PACKAGE =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     9
sig
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    10
  val add_nominal_datatype : bool -> string list -> (string list * bstring * mixfix *
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
    11
    (bstring * string list * mixfix) list) list -> theory -> theory
22433
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
    12
  type descr
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    13
  type nominal_datatype_info
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    14
  val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    15
  val get_nominal_datatype : theory -> string -> nominal_datatype_info option
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
    16
  val mk_perm: typ list -> term -> term -> term
22529
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
    17
  val perm_of_pair: term * term -> term
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
    18
  val mk_not_sym: thm list -> thm list
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
    19
  val perm_simproc: simproc
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    20
  val setup: theory -> theory
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    21
end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    22
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
    23
structure NominalPackage : NOMINAL_PACKAGE =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    24
struct
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    25
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
    26
val finite_emptyI = thm "finite.emptyI";
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    27
val finite_Diff = thm "finite_Diff";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    28
val finite_Un = thm "finite_Un";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    29
val Un_iff = thm "Un_iff";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    30
val In0_eq = thm "In0_eq";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    31
val In1_eq = thm "In1_eq";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    32
val In0_not_In1 = thm "In0_not_In1";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    33
val In1_not_In0 = thm "In1_not_In0";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    34
val Un_assoc = thm "Un_assoc";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    35
val Collect_disj_eq = thm "Collect_disj_eq";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    36
val empty_def = thm "empty_def";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    37
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    38
open DatatypeAux;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
    39
open NominalAtoms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    40
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    41
(** FIXME: DatatypePackage should export this function **)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    42
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    43
local
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    44
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    45
fun dt_recs (DtTFree _) = []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    46
  | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    47
  | dt_recs (DtRec i) = [i];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    48
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    49
fun dt_cases (descr: descr) (_, args, constrs) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    50
  let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    51
    fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i)));
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
    52
    val bnames = map the_bname (distinct op = (List.concat (map dt_recs args)));
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    53
  in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    54
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    55
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    56
fun induct_cases descr =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    57
  DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    58
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    59
fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    60
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    61
in
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    62
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    63
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    64
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    65
fun mk_case_names_exhausts descr new =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    66
  map (RuleCases.case_names o exhaust_cases descr o #1)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    67
    (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    68
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    69
end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    70
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    71
(* data kind 'HOL/nominal_datatypes' *)
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    72
22433
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
    73
type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list;
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
    74
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    75
type nominal_datatype_info =
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    76
  {index : int,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    77
   descr : descr,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    78
   sorts : (string * sort) list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    79
   rec_names : string list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    80
   rec_rewrites : thm list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    81
   induction : thm,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    82
   distinct : thm list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    83
   inject : thm list};
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    84
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    85
structure NominalDatatypesData = TheoryDataFun
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    86
(struct
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    87
  val name = "HOL/nominal_datatypes";
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    88
  type T = nominal_datatype_info Symtab.table;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    89
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    90
  val empty = Symtab.empty;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    91
  val copy = I;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    92
  val extend = I;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    93
  fun merge _ tabs : T = Symtab.merge (K true) tabs;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    94
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    95
  fun print sg tab =
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    96
    Pretty.writeln (Pretty.strs ("nominal datatypes:" ::
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    97
      map #1 (NameSpace.extern_table (Sign.type_space sg, tab))));
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    98
end);
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    99
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   100
val get_nominal_datatypes = NominalDatatypesData.get;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   101
val put_nominal_datatypes = NominalDatatypesData.put;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   102
val map_nominal_datatypes = NominalDatatypesData.map;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   103
val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   104
val setup = NominalDatatypesData.init;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   105
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   106
(**** make datatype info ****)
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   107
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   108
fun make_dt_info descr sorts induct reccomb_names rec_thms
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   109
    (((i, (_, (tname, _, _))), distinct), inject) =
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   110
  (tname,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   111
   {index = i,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   112
    descr = descr,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   113
    sorts = sorts,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   114
    rec_names = reccomb_names,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   115
    rec_rewrites = rec_thms,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   116
    induction = induct,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   117
    distinct = distinct,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   118
    inject = inject});
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
   119
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   120
(*******************************)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   121
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   122
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   123
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   124
fun read_typ sign ((Ts, sorts), str) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   125
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   126
    val T = Type.no_tvars (Sign.read_typ (sign, (AList.lookup op =)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   127
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   128
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   129
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   130
(** taken from HOL/Tools/datatype_aux.ML **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   131
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   132
fun indtac indrule indnames i st =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   133
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   134
    val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   135
    val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   136
      (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   137
    val getP = if can HOLogic.dest_imp (hd ts) then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   138
      (apfst SOME) o HOLogic.dest_imp else pair NONE;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   139
    fun abstr (t1, t2) = (case t1 of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   140
        NONE => (case filter (fn Free (s, _) => s mem indnames | _ => false)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   141
              (term_frees t2) of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   142
            [Free (s, T)] => absfree (s, T, t2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   143
          | _ => sys_error "indtac")
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   144
      | SOME (_ $ t') => Abs ("x", fastype_of t', abstract_over (t', t2)))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   145
    val cert = cterm_of (Thm.sign_of_thm st);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   146
    val Ps = map (cert o head_of o snd o getP) ts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   147
    val indrule' = cterm_instantiate (Ps ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   148
      (map (cert o abstr o getP) ts')) indrule
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   149
  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   150
    rtac indrule' i st
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   151
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   152
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   153
fun mk_subgoal i f st =
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   154
  let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   155
    val cg = List.nth (cprems_of st, i - 1);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   156
    val g = term_of cg;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   157
    val revcut_rl' = Thm.lift_rule cg revcut_rl;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   158
    val v = head_of (Logic.strip_assums_concl (hd (prems_of revcut_rl')));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   159
    val ps = Logic.strip_params g;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   160
    val cert = cterm_of (sign_of_thm st);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   161
  in
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   162
    compose_tac (false,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   163
      Thm.instantiate ([], [(cert v, cert (list_abs (ps,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   164
        f (rev ps) (Logic.strip_assums_hyp g) (Logic.strip_assums_concl g))))])
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   165
      revcut_rl', 2) i st
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   166
  end;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   167
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   168
(** simplification procedure for sorting permutations **)
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   169
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   170
val dj_cp = thm "dj_cp";
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   171
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   172
fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   173
      Type ("fun", [_, U])])) = (T, U);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   174
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   175
fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   176
  | permTs_of _ = [];
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   177
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   178
fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) =
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   179
      let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   180
        val (aT as Type (a, []), S) = dest_permT T;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   181
        val (bT as Type (b, []), _) = dest_permT U
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   182
      in if aT mem permTs_of u andalso aT <> bT then
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   183
          let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   184
            val a' = Sign.base_name a;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   185
            val b' = Sign.base_name b;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   186
            val cp = PureThy.get_thm thy (Name ("cp_" ^ a' ^ "_" ^ b' ^ "_inst"));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   187
            val dj = PureThy.get_thm thy (Name ("dj_" ^ b' ^ "_" ^ a'));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   188
            val dj_cp' = [cp, dj] MRS dj_cp;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   189
            val cert = SOME o cterm_of thy
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   190
          in
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   191
            SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   192
              [cert t, cert r, cert s] dj_cp'))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   193
          end
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   194
        else NONE
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   195
      end
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   196
  | perm_simproc' thy ss _ = NONE;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   197
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   198
val perm_simproc =
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   199
  Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \\<bullet> (pi2 \\<bullet> x)"] perm_simproc';
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   200
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   201
val allE_Nil = read_instantiate_sg (the_context()) [("x", "[]")] allE;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   202
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   203
val meta_spec = thm "meta_spec";
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   204
18582
4f4cc426b440 provide projections of induct_weak, induct_unsafe;
wenzelm
parents: 18579
diff changeset
   205
fun projections rule =
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19851
diff changeset
   206
  ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
18582
4f4cc426b440 provide projections of induct_weak, induct_unsafe;
wenzelm
parents: 18579
diff changeset
   207
  |> map (standard #> RuleCases.save rule);
4f4cc426b440 provide projections of induct_weak, induct_unsafe;
wenzelm
parents: 18579
diff changeset
   208
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
   209
val supp_prod = thm "supp_prod";
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   210
val fresh_prod = thm "fresh_prod";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   211
val supports_fresh = thm "supports_fresh";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   212
val supports_def = thm "Nominal.op supports_def";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   213
val fresh_def = thm "fresh_def";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   214
val supp_def = thm "supp_def";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   215
val rev_simps = thms "rev.simps";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
   216
val app_simps = thms "op @.simps";
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
   217
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   218
val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   219
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   220
fun mk_perm Ts t u =
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   221
  let
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   222
    val T = fastype_of1 (Ts, t);
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   223
    val U = fastype_of1 (Ts, u)
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   224
  in Const ("Nominal.perm", T --> U --> U) $ t $ u end;
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   225
22529
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   226
fun perm_of_pair (x, y) =
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   227
  let
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   228
    val T = fastype_of x;
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   229
    val pT = mk_permT T
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   230
  in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   231
    HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT)
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   232
  end;
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   233
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   234
fun mk_not_sym ths = maps (fn th => case prop_of th of
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   235
    _ $ (Const ("Not", _) $ _) => [th, th RS not_sym]
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   236
  | _ => [th]) ths;
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   237
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   238
fun gen_add_nominal_datatype prep_typ err flat_names new_type_names dts thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   239
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   240
    (* this theory is used just for parsing *)
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   241
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   242
    val tmp_thy = thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   243
      Theory.copy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   244
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   245
        (tname, length tvs, mx)) dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   246
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   247
    val sign = Theory.sign_of tmp_thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   248
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   249
    val atoms = atoms_of thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   250
    val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   251
    val cp_classes = List.concat (map (fn atom1 => map (fn atom2 =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   252
      Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   253
        Sign.base_name atom2)) atoms) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   254
    fun augment_sort S = S union classes;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   255
    val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   256
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   257
    fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   258
      let val (cargs', sorts') = Library.foldl (prep_typ sign) (([], sorts), cargs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   259
      in (constrs @ [(cname, cargs', mx)], sorts') end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   260
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   261
    fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   262
      let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   263
      in (dts @ [(tvs, tname, mx, constrs')], sorts') end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   264
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   265
    val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   266
    val sorts' = map (apsnd augment_sort) sorts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   267
    val tyvars = map #1 dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   268
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   269
    val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   270
    val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   271
      map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   272
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   273
    val ps = map (fn (_, n, _, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   274
      (Sign.full_name sign n, Sign.full_name sign (n ^ "_Rep"))) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   275
    val rps = map Library.swap ps;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   276
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   277
    fun replace_types (Type ("Nominal.ABS", [T, U])) =
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   278
          Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   279
      | replace_types (Type (s, Ts)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   280
          Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   281
      | replace_types T = T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   282
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   283
    val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn,
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   284
      map (fn (cname, cargs, mx) => (cname ^ "_Rep",
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   285
        map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   286
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   287
    val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   288
    val full_new_type_names' = map (Sign.full_name (sign_of thy)) new_type_names';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   289
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   290
    val ({induction, ...},thy1) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   291
      DatatypePackage.add_datatype_i err flat_names new_type_names' dts'' thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   292
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   293
    val SOME {descr, ...} = Symtab.lookup
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   294
      (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   295
    fun nth_dtyp i = typ_of_dtyp descr sorts' (DtRec i);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   296
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   297
    (**** define permutation functions ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   298
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   299
    val permT = mk_permT (TFree ("'x", HOLogic.typeS));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   300
    val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   301
    val perm_types = map (fn (i, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   302
      let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   303
      in permT --> T --> T end) descr;
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   304
    val perm_names = replicate (length new_type_names) "Nominal.perm" @
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   305
      DatatypeProp.indexify_names (map (fn i => Sign.full_name (sign_of thy1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   306
        ("perm_" ^ name_of_typ (nth_dtyp i)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   307
          (length new_type_names upto length descr - 1));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   308
    val perm_names_types = perm_names ~~ perm_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   309
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   310
    val perm_eqs = List.concat (map (fn (i, (_, _, constrs)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   311
      let val T = nth_dtyp i
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   312
      in map (fn (cname, dts) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   313
        let
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   314
          val Ts = map (typ_of_dtyp descr sorts') dts;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   315
          val names = DatatypeProp.make_tnames Ts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   316
          val args = map Free (names ~~ Ts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   317
          val c = Const (cname, Ts ---> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   318
          fun perm_arg (dt, x) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   319
            let val T = type_of x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   320
            in if is_rec_type dt then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   321
                let val (Us, _) = strip_type T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   322
                in list_abs (map (pair "x") Us,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   323
                  Const (List.nth (perm_names_types, body_index dt)) $ pi $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   324
                    list_comb (x, map (fn (i, U) =>
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   325
                      Const ("Nominal.perm", permT --> U --> U) $
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   326
                        (Const ("List.rev", permT --> permT) $ pi) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   327
                        Bound i) ((length Us - 1 downto 0) ~~ Us)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   328
                end
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   329
              else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   330
            end;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   331
        in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   332
          (("", HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   333
            (Const (List.nth (perm_names_types, i)) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   334
               Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   335
               list_comb (c, args),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   336
             list_comb (c, map perm_arg (dts ~~ args))))), [])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   337
        end) constrs
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   338
      end) descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   339
20179
a2f3f523c84b adaption to argument change in primrec_package
haftmann
parents: 20145
diff changeset
   340
    val (perm_simps, thy2) = thy1 |>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   341
      Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T, NoSyn))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   342
        (List.drop (perm_names_types, length new_type_names))) |>
19635
f7aa7d174343 unchecked definitions;
wenzelm
parents: 19494
diff changeset
   343
      PrimrecPackage.add_primrec_unchecked_i "" perm_eqs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   344
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   345
    (**** prove that permutation functions introduced by unfolding are ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   346
    (**** equivalent to already existing permutation functions         ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   347
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   348
    val _ = warning ("length descr: " ^ string_of_int (length descr));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   349
    val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   350
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   351
    val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   352
    val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   353
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   354
    val unfolded_perm_eq_thms =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   355
      if length descr = length new_type_names then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   356
      else map standard (List.drop (split_conj_thm
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   357
        (Goal.prove_global thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   358
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   359
            (map (fn (c as (s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   360
               let val [T1, T2] = binder_types T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   361
               in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   362
                 Const ("Nominal.perm", T) $ pi $ Free (x, T2))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   363
               end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   364
             (perm_names_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   365
          (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   366
            ALLGOALS (asm_full_simp_tac
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   367
              (simpset_of thy2 addsimps [perm_fun_def]))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   368
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   369
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   370
    (**** prove [] \<bullet> t = t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   371
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   372
    val _ = warning "perm_empty_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   373
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   374
    val perm_empty_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   375
      let val permT = mk_permT (Type (a, []))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   376
      in map standard (List.take (split_conj_thm
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   377
        (Goal.prove_global thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   378
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   379
            (map (fn ((s, T), x) => HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   380
                (Const (s, permT --> T --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   381
                   Const ("List.list.Nil", permT) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   382
                 Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   383
             (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   384
              map body_type perm_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   385
          (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   386
            ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   387
        length new_type_names))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   388
      end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   389
      atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   390
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   391
    (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   392
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   393
    val _ = warning "perm_append_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   394
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   395
    (*FIXME: these should be looked up statically*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   396
    val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   397
    val pt2 = PureThy.get_thm thy2 (Name "pt2");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   398
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   399
    val perm_append_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   400
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   401
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   402
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   403
        val pi2 = Free ("pi2", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   404
        val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   405
        val pt2' = pt_inst RS pt2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   406
        val pt2_ax = PureThy.get_thm thy2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   407
          (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   408
      in List.take (map standard (split_conj_thm
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   409
        (Goal.prove_global thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   410
             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   411
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   412
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   413
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   414
                      (perm $ (Const ("List.op @", permT --> permT --> permT) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   415
                         pi1 $ pi2) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   416
                       perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   417
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   418
                  (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   419
                   map body_type perm_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   420
           (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   421
              ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   422
         length new_type_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   423
      end) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   424
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   425
    (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   426
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   427
    val _ = warning "perm_eq_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   428
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   429
    val pt3 = PureThy.get_thm thy2 (Name "pt3");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   430
    val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   431
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   432
    val perm_eq_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   433
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   434
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   435
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   436
        val pi2 = Free ("pi2", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   437
        (*FIXME: not robust - better access these theorems using NominalData?*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   438
        val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   439
        val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   440
        val pt3' = pt_inst RS pt3;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   441
        val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   442
        val pt3_ax = PureThy.get_thm thy2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   443
          (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   444
      in List.take (map standard (split_conj_thm
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   445
        (Goal.prove_global thy2 [] [] (Logic.mk_implies
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   446
             (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   447
                permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   448
              HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   449
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   450
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   451
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   452
                      (perm $ pi1 $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   453
                       perm $ pi2 $ Free (x, T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   454
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   455
                  (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   456
                   map body_type perm_types ~~ perm_indnames)))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   457
           (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   458
              ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   459
         length new_type_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   460
      end) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   461
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   462
    (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   463
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   464
    val cp1 = PureThy.get_thm thy2 (Name "cp1");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   465
    val dj_cp = PureThy.get_thm thy2 (Name "dj_cp");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   466
    val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   467
    val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   468
    val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   469
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   470
    fun composition_instance name1 name2 thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   471
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   472
        val name1' = Sign.base_name name1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   473
        val name2' = Sign.base_name name2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   474
        val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   475
        val permT1 = mk_permT (Type (name1, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   476
        val permT2 = mk_permT (Type (name2, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   477
        val augment = map_type_tfree
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   478
          (fn (x, S) => TFree (x, cp_class :: S));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   479
        val Ts = map (augment o body_type) perm_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   480
        val cp_inst = PureThy.get_thm thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   481
          (Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   482
        val simps = simpset_of thy addsimps (perm_fun_def ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   483
          (if name1 <> name2 then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   484
             let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1'))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   485
             in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   486
           else
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   487
             let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   488
               val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   489
               val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst"))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   490
             in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   491
               [cp_inst RS cp1 RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   492
                at_inst RS (pt_inst RS pt_perm_compose) RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   493
                at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   494
            end))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   495
        val thms = split_conj_thm (Goal.prove_global thy [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   496
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   497
              (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   498
                  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   499
                    val pi1 = Free ("pi1", permT1);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   500
                    val pi2 = Free ("pi2", permT2);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   501
                    val perm1 = Const (s, permT1 --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   502
                    val perm2 = Const (s, permT2 --> T --> T);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   503
                    val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   504
                  in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   505
                    (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   506
                     perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   507
                  end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   508
                (perm_names ~~ Ts ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   509
          (fn _ => EVERY [indtac induction perm_indnames 1,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   510
             ALLGOALS (asm_full_simp_tac simps)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   511
      in
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19251
diff changeset
   512
        foldl (fn ((s, tvs), thy) => AxClass.prove_arity
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   513
            (s, replicate (length tvs) (cp_class :: classes), [cp_class])
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   514
            (ClassPackage.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   515
          thy (full_new_type_names' ~~ tyvars)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   516
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   517
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   518
    val (perm_thmss,thy3) = thy2 |>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   519
      fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   520
      curry (Library.foldr (fn ((i, (tyname, args, _)), thy) =>
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19251
diff changeset
   521
        AxClass.prove_arity (tyname, replicate (length args) classes, classes)
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   522
        (ClassPackage.intro_classes_tac [] THEN REPEAT (EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   523
           [resolve_tac perm_empty_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   524
            resolve_tac perm_append_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   525
            resolve_tac perm_eq_thms 1, assume_tac 1])) thy))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   526
        (List.take (descr, length new_type_names)) |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   527
      PureThy.add_thmss
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   528
        [((space_implode "_" new_type_names ^ "_unfolded_perm_eq",
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   529
          unfolded_perm_eq_thms), [Simplifier.simp_add]),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   530
         ((space_implode "_" new_type_names ^ "_perm_empty",
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   531
          perm_empty_thms), [Simplifier.simp_add]),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   532
         ((space_implode "_" new_type_names ^ "_perm_append",
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   533
          perm_append_thms), [Simplifier.simp_add]),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   534
         ((space_implode "_" new_type_names ^ "_perm_eq",
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   535
          perm_eq_thms), [Simplifier.simp_add])];
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   536
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   537
    (**** Define representing sets ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   538
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   539
    val _ = warning "representing sets";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   540
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   541
    val rep_set_names = DatatypeProp.indexify_names
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   542
      (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   543
    val big_rep_name =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   544
      space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   545
        (fn (i, ("Nominal.noption", _, _)) => NONE
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   546
          | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   547
    val _ = warning ("big_rep_name: " ^ big_rep_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   548
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   549
    fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   550
          (case AList.lookup op = descr i of
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   551
             SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   552
               apfst (cons dt) (strip_option dt')
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   553
           | _ => ([], dtf))
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   554
      | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   555
          apfst (cons dt) (strip_option dt')
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   556
      | strip_option dt = ([], dt);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   557
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   558
    val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts')
18280
45e139675daf Corrected atom class constraints in strong induction rule.
berghofe
parents: 18261
diff changeset
   559
      (List.concat (map (fn (_, (_, _, cs)) => List.concat
45e139675daf Corrected atom class constraints in strong induction rule.
berghofe
parents: 18261
diff changeset
   560
        (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
45e139675daf Corrected atom class constraints in strong induction rule.
berghofe
parents: 18261
diff changeset
   561
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   562
    fun make_intr s T (cname, cargs) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   563
      let
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   564
        fun mk_prem (dt, (j, j', prems, ts)) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   565
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   566
            val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   567
            val (dts', dt'') = strip_dtyp dt';
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   568
            val Ts = map (typ_of_dtyp descr sorts') dts;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   569
            val Us = map (typ_of_dtyp descr sorts') dts';
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   570
            val T = typ_of_dtyp descr sorts' dt'';
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   571
            val free = mk_Free "x" (Us ---> T) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   572
            val free' = app_bnds free (length Us);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   573
            fun mk_abs_fun (T, (i, t)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   574
              let val U = fastype_of t
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   575
              in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   576
                Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   577
              end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   578
          in (j + 1, j' + length Ts,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   579
            case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   580
                DtRec k => list_all (map (pair "x") Us,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   581
                  HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k),
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   582
                    T --> HOLogic.boolT) $ free')) :: prems
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   583
              | _ => prems,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   584
            snd (foldr mk_abs_fun (j', free) Ts) :: ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   585
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   586
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   587
        val (_, _, prems, ts) = foldr mk_prem (1, 1, [], []) cargs;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   588
        val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   589
          list_comb (Const (cname, map fastype_of ts ---> T), ts))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   590
      in Logic.list_implies (prems, concl)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   591
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   592
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   593
    val (intr_ts, (rep_set_names', recTs')) =
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   594
      apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   595
        (fn ((_, ("Nominal.noption", _, _)), _) => NONE
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   596
          | ((i, (_, _, constrs)), rep_set_name) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   597
              let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   598
              in SOME (map (make_intr rep_set_name T) constrs,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   599
                (rep_set_name, T))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   600
              end)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   601
                (descr ~~ rep_set_names))));
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   602
    val rep_set_names'' = map (Sign.full_name thy3) rep_set_names';
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   603
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   604
    val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   605
      setmp InductivePackage.quiet_mode false
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   606
        (TheoryTarget.init NONE #>
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   607
         InductivePackage.add_inductive_i false big_rep_name false true false
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   608
           (map (fn (s, T) => (s, SOME (T --> HOLogic.boolT), NoSyn))
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   609
              (rep_set_names' ~~ recTs'))
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   610
           [] (map (fn x => (("", []), x)) intr_ts) [] #>
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   611
         apsnd (ProofContext.theory_of o LocalTheory.exit)) thy3;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   612
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   613
    (**** Prove that representing set is closed under permutation ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   614
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   615
    val _ = warning "proving closure under permutation...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   616
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   617
    val perm_indnames' = List.mapPartial
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   618
      (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   619
      (perm_indnames ~~ descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   620
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   621
    fun mk_perm_closed name = map (fn th => standard (th RS mp))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   622
      (List.take (split_conj_thm (Goal.prove_global thy4 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   623
        (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   624
           (fn ((s, T), x) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   625
              let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   626
                val T = map_type_tfree
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   627
                  (fn (s, cs) => TFree (s, cs union cp_classes)) T;
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   628
                val S = Const (s, T --> HOLogic.boolT);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   629
                val permT = mk_permT (Type (name, []))
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   630
              in HOLogic.mk_imp (S $ Free (x, T),
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   631
                S $ (Const ("Nominal.perm", permT --> T --> T) $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   632
                  Free ("pi", permT) $ Free (x, T)))
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   633
              end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   634
        (fn _ => EVERY (* CU: added perm_fun_def in the final tactic in order to deal with funs *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   635
           [indtac rep_induct [] 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   636
            ALLGOALS (simp_tac (simpset_of thy4 addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   637
              (symmetric perm_fun_def :: PureThy.get_thms thy4 (Name ("abs_perm"))))),
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   638
            ALLGOALS (resolve_tac rep_intrs
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   639
               THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   640
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   641
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   642
    (* FIXME: theorems are stored in database for testing only *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   643
    val perm_closed_thmss = map mk_perm_closed atoms;
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20411
diff changeset
   644
    val (_, thy5) = PureThy.add_thmss [(("perm_closed", List.concat perm_closed_thmss), [])] thy4;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   645
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   646
    (**** typedef ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   647
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   648
    val _ = warning "defining type...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   649
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   650
    val (typedefs, thy6) =
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20411
diff changeset
   651
      thy5
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   652
      |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   653
        setmp TypedefPackage.quiet_mode true
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   654
          (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx)
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   655
            (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   656
               Const (cname, U --> HOLogic.boolT)) NONE
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   657
            (rtac exI 1 THEN rtac CollectI 1 THEN
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   658
              QUIET_BREADTH_FIRST (has_fewer_prems 1)
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20411
diff changeset
   659
              (resolve_tac rep_intrs 1))) thy |> (fn ((_, r), thy) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   660
        let
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
   661
          val permT = mk_permT (TFree (Name.variant tvs "'a", HOLogic.typeS));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   662
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   663
          val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   664
          val T = Type (Sign.intern_type thy name, tvs');
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   665
        in apfst (pair r o hd)
19635
f7aa7d174343 unchecked definitions;
wenzelm
parents: 19494
diff changeset
   666
          (PureThy.add_defs_unchecked_i true [(("prm_" ^ name ^ "_def", Logic.mk_equals
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   667
            (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   668
             Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   669
               (Const ("Nominal.perm", permT --> U --> U) $ pi $
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   670
                 (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   671
                   Free ("x", T))))), [])] thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   672
        end))
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   673
          (types_syntax ~~ tyvars ~~
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   674
            List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   675
            new_type_names);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   676
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   677
    val perm_defs = map snd typedefs;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   678
    val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   679
    val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   680
    val Rep_thms = map (collect_simp o #Rep o fst) typedefs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   681
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   682
    val big_name = space_implode "_" new_type_names;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   683
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   684
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   685
    (** prove that new types are in class pt_<name> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   686
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   687
    val _ = warning "prove that new types are in class pt_<name> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   688
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   689
    fun pt_instance ((class, atom), perm_closed_thms) =
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   690
      fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   691
        perm_def), name), tvs), perm_closed) => fn thy =>
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19251
diff changeset
   692
          AxClass.prove_arity
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   693
            (Sign.intern_type thy name,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   694
              replicate (length tvs) (classes @ cp_classes), [class])
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   695
            (EVERY [ClassPackage.intro_classes_tac [],
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   696
              rewrite_goals_tac [perm_def],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   697
              asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   698
              asm_full_simp_tac (simpset_of thy addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   699
                [Rep RS perm_closed RS Abs_inverse]) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   700
              asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   701
                (Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   702
        (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   703
           new_type_names ~~ tyvars ~~ perm_closed_thms);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   704
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   705
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   706
    (** prove that new types are in class cp_<name1>_<name2> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   707
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   708
    val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   709
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   710
    fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   711
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   712
        val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   713
        val class = Sign.intern_class thy name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   714
        val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   715
      in fold (fn ((((((Abs_inverse, Rep),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   716
        perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19251
diff changeset
   717
          AxClass.prove_arity
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   718
            (Sign.intern_type thy name,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   719
              replicate (length tvs) (classes @ cp_classes), [class])
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   720
            (EVERY [ClassPackage.intro_classes_tac [],
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   721
              rewrite_goals_tac [perm_def],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   722
              asm_full_simp_tac (simpset_of thy addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   723
                ((Rep RS perm_closed1 RS Abs_inverse) ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   724
                 (if atom1 = atom2 then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   725
                  else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   726
              cong_tac 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   727
              rtac refl 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   728
              rtac cp1' 1]) thy)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   729
        (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   730
           tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   731
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   732
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   733
    val thy7 = fold (fn x => fn thy => thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   734
      pt_instance x |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   735
      fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   736
        (classes ~~ atoms ~~ perm_closed_thmss) thy6;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   737
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   738
    (**** constructors ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   739
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   740
    fun mk_abs_fun (x, t) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   741
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   742
        val T = fastype_of x;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   743
        val U = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   744
      in
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   745
        Const ("Nominal.abs_fun", T --> U --> T -->
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   746
          Type ("Nominal.noption", [U])) $ x $ t
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   747
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   748
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   749
    val (ty_idxs, _) = foldl
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   750
      (fn ((i, ("Nominal.noption", _, _)), p) => p
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   751
        | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   752
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   753
    fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   754
      | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   755
      | reindex dt = dt;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   756
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   757
    fun strip_suffix i s = implode (List.take (explode s, size s - i));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   758
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   759
    (** strips the "_Rep" in type names *)
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   760
    fun strip_nth_name i s =
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21669
diff changeset
   761
      let val xs = NameSpace.explode s;
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21669
diff changeset
   762
      in NameSpace.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   763
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   764
    val (descr'', ndescr) = ListPair.unzip (List.mapPartial
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   765
      (fn (i, ("Nominal.noption", _, _)) => NONE
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   766
        | (i, (s, dts, constrs)) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   767
             let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   768
               val SOME index = AList.lookup op = ty_idxs i;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   769
               val (constrs1, constrs2) = ListPair.unzip
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   770
                 (map (fn (cname, cargs) => apfst (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   771
                   (foldl_map (fn (dts, dt) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   772
                     let val (dts', dt') = strip_option dt
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   773
                     in (dts @ dts' @ [reindex dt'], (length dts, length dts')) end)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   774
                       ([], cargs))) constrs)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   775
             in SOME ((index, (strip_nth_name 1 s,  map reindex dts, constrs1)),
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   776
               (index, constrs2))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   777
             end) descr);
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   778
19489
4441b637871b SplitAt -> chop
berghofe
parents: 19403
diff changeset
   779
    val (descr1, descr2) = chop (length new_type_names) descr'';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   780
    val descr' = [descr1, descr2];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   781
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
   782
    fun partition_cargs idxs xs = map (fn (i, j) =>
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
   783
      (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs;
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
   784
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   785
    val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   786
      map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   787
        (constrs ~~ idxss)))) (descr'' ~~ ndescr);
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   788
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   789
    fun nth_dtyp' i = typ_of_dtyp descr'' sorts' (DtRec i);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   790
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   791
    val rep_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   792
      Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   793
    val abs_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   794
      Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   795
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   796
    val recTs = get_rec_types descr'' sorts';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   797
    val newTs' = Library.take (length new_type_names, recTs');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   798
    val newTs = Library.take (length new_type_names, recTs);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   799
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   800
    val full_new_type_names = map (Sign.full_name (sign_of thy)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   801
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   802
    fun make_constr_def tname T T' ((thy, defs, eqns),
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   803
        (((cname_rep, _), (cname, cargs)), (cname', mx))) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   804
      let
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   805
        fun constr_arg ((dts, dt), (j, l_args, r_args)) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   806
          let
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   807
            val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts' dt) i)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   808
              (dts ~~ (j upto j + length dts - 1))
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   809
            val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   810
          in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   811
            (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   812
             xs @ x :: l_args,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   813
             foldr mk_abs_fun
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   814
               (case dt of
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   815
                  DtRec k => if k < length new_type_names then
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   816
                      Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts' dt -->
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   817
                        typ_of_dtyp descr sorts' dt) $ x
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   818
                    else error "nested recursion not (yet) supported"
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   819
                | _ => x) xs :: r_args)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   820
          end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   821
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   822
        val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   823
        val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   824
        val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   825
        val constrT = map fastype_of l_args ---> T;
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   826
        val lhs = list_comb (Const (cname, constrT), l_args);
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   827
        val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   828
        val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   829
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   830
          (Const (rep_name, T --> T') $ lhs, rhs));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   831
        val def_name = (Sign.base_name cname) ^ "_def";
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   832
        val ([def_thm], thy') = thy |>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   833
          Theory.add_consts_i [(cname', constrT, mx)] |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   834
          (PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   835
      in (thy', defs @ [def_thm], eqns @ [eqn]) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   836
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   837
    fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   838
        (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   839
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   840
        val rep_const = cterm_of thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   841
          (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   842
        val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   843
        val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   844
          ((Theory.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   845
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   846
        (parent_path flat_names thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   847
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   848
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   849
    val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   850
      ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   851
        List.take (pdescr, length new_type_names) ~~
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   852
        new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   853
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   854
    val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   855
    val rep_inject_thms = map (#Rep_inject o fst) typedefs
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   856
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   857
    (* prove theorem  Rep_i (Constr_j ...) = Constr'_j ...  *)
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   858
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   859
    fun prove_constr_rep_thm eqn =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   860
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   861
        val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   862
        val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   863
      in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   864
        [resolve_tac inj_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   865
         rewrite_goals_tac rewrites,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   866
         rtac refl 3,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   867
         resolve_tac rep_intrs 2,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   868
         REPEAT (resolve_tac Rep_thms 1)])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   869
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   871
    val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   872
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   873
    (* prove theorem  pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   874
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   875
    fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   876
      let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   877
        val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   878
        val Type ("fun", [T, U]) = fastype_of Rep;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   879
        val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   880
        val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   881
      in
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   882
        Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   883
            (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   884
             Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x))))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   885
          (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   886
            perm_closed_thms @ Rep_thms)) 1)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   887
      end) Rep_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   888
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   889
    val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   890
      (atoms ~~ perm_closed_thmss));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   891
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   892
    (* prove distinctness theorems *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   893
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   894
    val distinct_props = setmp DatatypeProp.dtK 1000
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   895
      (DatatypeProp.make_distincts new_type_names descr' sorts') thy8;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   896
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   897
    val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   898
      dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   899
        (constr_rep_thmss ~~ dist_lemmas);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   900
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   901
    fun prove_distinct_thms (_, []) = []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   902
      | prove_distinct_thms (p as (rep_thms, dist_lemma), t::ts) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   903
          let
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   904
            val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   905
              simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   906
          in dist_thm::(standard (dist_thm RS not_sym))::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   907
            (prove_distinct_thms (p, ts))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   908
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   909
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   910
    val distinct_thms = map prove_distinct_thms
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   911
      (constr_rep_thmss ~~ dist_lemmas ~~ distinct_props);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   912
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   913
    (** prove equations for permutation functions **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   914
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   915
    val abs_perm = PureThy.get_thms thy8 (Name "abs_perm"); (* FIXME *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   916
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   917
    val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   918
      let val T = nth_dtyp' i
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   919
      in List.concat (map (fn (atom, perm_closed_thms) =>
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   920
          map (fn ((cname, dts), constr_rep_thm) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   921
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   922
          val cname = Sign.intern_const thy8
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   923
            (NameSpace.append tname (Sign.base_name cname));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   924
          val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   925
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   926
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   927
          fun perm t =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   928
            let val T = fastype_of t
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
   929
            in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   930
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   931
          fun constr_arg ((dts, dt), (j, l_args, r_args)) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   932
            let
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   933
              val Ts = map (typ_of_dtyp descr'' sorts') dts;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   934
              val xs = map (fn (T, i) => mk_Free "x" T i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   935
                (Ts ~~ (j upto j + length dts - 1))
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   936
              val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   937
            in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   938
              (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   939
               xs @ x :: l_args,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   940
               map perm (xs @ [x]) @ r_args)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   941
            end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   942
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   943
          val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   944
          val c = Const (cname, map fastype_of l_args ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   945
        in
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   946
          Goal.prove_global thy8 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   947
            (HOLogic.mk_Trueprop (HOLogic.mk_eq
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   948
              (perm (list_comb (c, l_args)), list_comb (c, r_args))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   949
            (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   950
              [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   951
               simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   952
                 constr_defs @ perm_closed_thms)) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   953
               TRY (simp_tac (HOL_basic_ss addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   954
                 (symmetric perm_fun_def :: abs_perm)) 1),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   955
               TRY (simp_tac (HOL_basic_ss addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   956
                 (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   957
                    perm_closed_thms)) 1)])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   958
        end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   959
      end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   960
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   961
    (** prove injectivity of constructors **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   962
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   963
    val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   964
    val alpha = PureThy.get_thms thy8 (Name "alpha");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   965
    val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   966
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   967
    val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   968
      let val T = nth_dtyp' i
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   969
      in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   970
        if null dts then NONE else SOME
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   971
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   972
          val cname = Sign.intern_const thy8
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   973
            (NameSpace.append tname (Sign.base_name cname));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   974
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   975
          fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   976
            let
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   977
              val Ts_idx = map (typ_of_dtyp descr'' sorts') dts ~~ (j upto j + length dts - 1);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   978
              val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   979
              val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   980
              val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts);
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   981
              val y = mk_Free "y" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   982
            in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   983
              (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   984
               xs @ (x :: args1), ys @ (y :: args2),
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   985
               HOLogic.mk_eq
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   986
                 (foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   987
            end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   988
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   989
          val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   990
          val Ts = map fastype_of args1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   991
          val c = Const (cname, Ts ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   992
        in
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   993
          Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   994
              (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   995
               foldr1 HOLogic.mk_conj eqs)))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   996
            (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   997
               [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   998
                  rep_inject_thms')) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   999
                TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1000
                  alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
17874
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
  1001
                  perm_rep_perm_thms)) 1),
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
  1002
                TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def ::
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1003
                  expand_fun_eq :: rep_inject_thms @ perm_rep_perm_thms)) 1)])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1004
        end) (constrs ~~ constr_rep_thms)
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1005
      end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1006
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1007
    (** equations for support and freshness **)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1008
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1009
    val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1010
      (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1011
      let val T = nth_dtyp' i
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1012
      in List.concat (map (fn (cname, dts) => map (fn atom =>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1013
        let
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1014
          val cname = Sign.intern_const thy8
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1015
            (NameSpace.append tname (Sign.base_name cname));
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1016
          val atomT = Type (atom, []);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1017
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1018
          fun process_constr ((dts, dt), (j, args1, args2)) =
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1019
            let
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1020
              val Ts_idx = map (typ_of_dtyp descr'' sorts') dts ~~ (j upto j + length dts - 1);
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1021
              val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1022
              val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts)
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
  1023
            in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
  1024
              (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
  1025
               xs @ (x :: args1), foldr mk_abs_fun x xs :: args2)
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1026
            end;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1027
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1028
          val (_, args1, args2) = foldr process_constr (1, [], []) dts;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1029
          val Ts = map fastype_of args1;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1030
          val c = list_comb (Const (cname, Ts ---> T), args1);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1031
          fun supp t =
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
  1032
            Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1033
          fun fresh t =
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
  1034
            Const ("Nominal.fresh", atomT --> fastype_of t --> HOLogic.boolT) $
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1035
              Free ("a", atomT) $ t;
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1036
          val supp_thm = Goal.prove_global thy8 [] []
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1037
              (HOLogic.mk_Trueprop (HOLogic.mk_eq
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1038
                (supp c,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1039
                 if null dts then Const ("{}", HOLogic.mk_setT atomT)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1040
                 else foldr1 (HOLogic.mk_binop "op Un") (map supp args2))))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1041
            (fn _ =>
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1042
              simp_tac (HOL_basic_ss addsimps (supp_def ::
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1043
                 Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1044
                 symmetric empty_def :: finite_emptyI :: simp_thms @
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1045
                 abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1046
        in
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1047
          (supp_thm,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1048
           Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1049
              (fresh c,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1050
               if null dts then HOLogic.true_const
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1051
               else foldr1 HOLogic.mk_conj (map fresh args2))))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1052
             (fn _ =>
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1053
               simp_tac (simpset_of thy8 addsimps [fresh_def, supp_thm]) 1))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1054
        end) atoms) constrs)
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1055
      end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1056
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1057
    (**** weak induction theorem ****)
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1058
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1059
    fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1060
      let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1061
        val Rep_t = Const (List.nth (rep_names, i), T --> U) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1062
          mk_Free "x" T i;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1063
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1064
        val Abs_t =  Const (List.nth (abs_names, i), U --> T)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1065
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1066
      in (prems @ [HOLogic.imp $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1067
            (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1068
              (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1069
          concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1070
      end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1071
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1072
    val (indrule_lemma_prems, indrule_lemma_concls) =
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1073
      Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1074
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1075
    val indrule_lemma = Goal.prove_global thy8 [] []
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1076
      (Logic.mk_implies
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1077
        (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1078
         HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1079
           [REPEAT (etac conjE 1),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1080
            REPEAT (EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1081
              [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1082
               etac mp 1, resolve_tac Rep_thms 1])]);
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1083
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1084
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1085
    val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1086
      map (Free o apfst fst o dest_Var) Ps;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1087
    val indrule_lemma' = cterm_instantiate
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1088
      (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1089
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1090
    val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1091
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1092
    val dt_induct_prop = DatatypeProp.make_ind descr' sorts';
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1093
    val dt_induct = Goal.prove_global thy8 []
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1094
      (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1095
      (fn prems => EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1096
        [rtac indrule_lemma' 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1097
         (DatatypeAux.indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1098
         EVERY (map (fn (prem, r) => (EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1099
           [REPEAT (eresolve_tac Abs_inverse_thms' 1),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1100
            simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1101
            DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1102
                (prems ~~ constr_defs))]);
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1103
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1104
    val case_names_induct = mk_case_names_induct descr'';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1105
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1106
    (**** prove that new datatypes have finite support ****)
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1107
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1108
    val _ = warning "proving finite support for the new datatype";
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1109
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1110
    val indnames = DatatypeProp.make_tnames recTs;
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1111
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1112
    val abs_supp = PureThy.get_thms thy8 (Name "abs_supp");
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
  1113
    val supp_atm = PureThy.get_thms thy8 (Name "supp_atm");
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1114
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1115
    val finite_supp_thms = map (fn atom =>
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1116
      let val atomT = Type (atom, [])
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1117
      in map standard (List.take
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1118
        (split_conj_thm (Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1119
           (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1120
             Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1121
               (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1122
               (indnames ~~ recTs))))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1123
           (fn _ => indtac dt_induct indnames 1 THEN
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1124
            ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
  1125
              (abs_supp @ supp_atm @
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1126
               PureThy.get_thms thy8 (Name ("fs_" ^ Sign.base_name atom ^ "1")) @
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1127
               List.concat supp_thms))))),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1128
         length new_type_names))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1129
      end) atoms;
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1130
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
  1131
    val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1132
22245
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1133
	(* Function to add both the simp and eqvt attributes *)
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1134
        (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1135
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1136
    val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1137
 
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1138
    val (_, thy9) = thy8 |>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1139
      Theory.add_path big_name |>
22543
9460a4cf0acc Renamed induct(s)_weak to weak_induct(s) in order to unify
berghofe
parents: 22529
diff changeset
  1140
      PureThy.add_thms [(("weak_induct", dt_induct), [case_names_induct])] ||>>
9460a4cf0acc Renamed induct(s)_weak to weak_induct(s) in order to unify
berghofe
parents: 22529
diff changeset
  1141
      PureThy.add_thmss [(("weak_inducts", projections dt_induct), [case_names_induct])] ||>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1142
      Theory.parent_path ||>>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1143
      DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1144
      DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
22231
f76f187c91f9 added an infrastructure that allows the user to declare lemmas to be equivariance lemmas; the intention is to use these lemmas in automated tools but also can be employed by the user
urbanc
parents: 21858
diff changeset
  1145
      DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1146
      DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1147
      DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1148
      DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1149
      fold (fn (atom, ths) => fn thy =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1150
        let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom)
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19251
diff changeset
  1151
        in fold (fn T => AxClass.prove_arity
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1152
            (fst (dest_Type T),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1153
              replicate (length sorts) [class], [class])
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
  1154
            (ClassPackage.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1155
        end) (atoms ~~ finite_supp_thms);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1156
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1157
    (**** strong induction theorem ****)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1158
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1159
    val pnames = if length descr'' = 1 then ["P"]
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1160
      else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
18245
65e60434b3c2 Fixed problem with strong induction theorem for datatypes containing
berghofe
parents: 18142
diff changeset
  1161
    val ind_sort = if null dt_atomTs then HOLogic.typeS
19649
c887656778bc Sign.certify_sort;
wenzelm
parents: 19635
diff changeset
  1162
      else Sign.certify_sort thy9 (map (fn T => Sign.intern_class thy9 ("fs_" ^
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1163
        Sign.base_name (fst (dest_Type T)))) dt_atomTs);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1164
    val fsT = TFree ("'n", ind_sort);
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1165
    val fsT' = TFree ("'n", HOLogic.typeS);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1166
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1167
    val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1168
      (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1169
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1170
    fun make_pred fsT i T =
18302
577e5d19b33c Changed order of predicate arguments and quantifiers in strong induction rule.
berghofe
parents: 18280
diff changeset
  1171
      Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1172
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1173
    fun mk_fresh1 xs [] = []
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1174
      | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1175
            (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1176
              (filter (fn (_, U) => T = U) (rev xs)) @
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1177
          mk_fresh1 (y :: xs) ys;
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1178
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1179
    fun mk_fresh2 xss [] = []
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1180
      | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) =>
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1181
            map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1182
              (Const ("Nominal.fresh", T --> U --> HOLogic.boolT) $ Free y $ Free x))
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1183
                (rev xss @ yss)) ys) @
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1184
          mk_fresh2 (p :: xss) yss;
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1185
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1186
    fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1187
      let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1188
        val recs = List.filter is_rec_type cargs;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1189
        val Ts = map (typ_of_dtyp descr'' sorts') cargs;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1190
        val recTs' = map (typ_of_dtyp descr'' sorts') recs;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
  1191
        val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1192
        val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1193
        val frees = tnames ~~ Ts;
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1194
        val frees' = partition_cargs idxs frees;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
  1195
        val z = (Name.variant tnames "z", fsT);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1196
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1197
        fun mk_prem ((dt, s), T) =
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1198
          let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1199
            val (Us, U) = strip_type T;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1200
            val l = length Us
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1201
          in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1202
            (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l))
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1203
          end;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1204
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1205
        val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1206
        val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1207
            (f T (Free p) (Free z))) (List.concat (map fst frees')) @
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1208
          mk_fresh1 [] (List.concat (map fst frees')) @
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1209
          mk_fresh2 [] frees'
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1210
18302
577e5d19b33c Changed order of predicate arguments and quantifiers in strong induction rule.
berghofe
parents: 18280
diff changeset
  1211
      in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems,
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1212
        HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
18302
577e5d19b33c Changed order of predicate arguments and quantifiers in strong induction rule.
berghofe
parents: 18280
diff changeset
  1213
          list_comb (Const (cname, Ts ---> T), map Free frees))))
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1214
      end;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1215
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1216
    val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1217
      map (make_ind_prem fsT (fn T => fn t => fn u =>
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
  1218
        Const ("Nominal.fresh", T --> fsT --> HOLogic.boolT) $ t $ u) i T)
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1219
          (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1220
    val tnames = DatatypeProp.make_tnames recTs;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
  1221
    val zs = Name.variant_list tnames (replicate (length descr'') "z");
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1222
    val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1223
      (map (fn ((((i, _), T), tname), z) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1224
        make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1225
        (descr'' ~~ recTs ~~ tnames ~~ zs)));
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1226
    val induct = Logic.list_implies (ind_prems, ind_concl);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1227
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1228
    val ind_prems' =
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1229
      map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')],
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1230
        HOLogic.mk_Trueprop (Const ("Finite_Set.finite", body_type T -->
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1231
          HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1232
      List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1233
        map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1234
          HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1235
            (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1236
    val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1237
      (map (fn ((((i, _), T), tname), z) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1238
        make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1239
        (descr'' ~~ recTs ~~ tnames ~~ zs)));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1240
    val induct' = Logic.list_implies (ind_prems', ind_concl');
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1241
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1242
    val aux_ind_vars =
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1243
      (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1244
       map mk_permT dt_atomTs) @ [("z", fsT')];
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1245
    val aux_ind_Ts = rev (map snd aux_ind_vars);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1246
    val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1247
      (map (fn (((i, _), T), tname) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1248
        HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1249
          fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1250
            (Free (tname, T))))
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1251
        (descr'' ~~ recTs ~~ tnames)));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1252
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1253
    fun mk_ind_perm i k p l vs j =
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1254
      let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1255
        val n = length vs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1256
        val Ts = map snd vs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1257
        val T = List.nth (Ts, i - j);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1258
        val pT = NominalAtoms.mk_permT T
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1259
      in
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1260
        Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1261
          (HOLogic.pair_const T T $ Bound (l - j) $ fold_rev (mk_perm Ts)
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1262
            (map (mk_ind_perm i k p l vs) (j - 1 downto 0) @
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1263
             map Bound (n - k - 1 downto n - k - p))
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1264
            (Bound (i - j))) $
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1265
          Const ("List.list.Nil", pT)
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1266
      end;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1267
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1268
    fun mk_fresh i i' j k p l is vs _ _ =
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1269
      let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1270
        val n = length vs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1271
        val Ts = map snd vs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1272
        val T = List.nth (Ts, n - i - 1 - j);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1273
        val f = the (AList.lookup op = fresh_fs T);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1274
        val U = List.nth (Ts, n - i' - 1);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1275
        val S = HOLogic.mk_setT T;
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1276
        val prms' = map Bound (n - k downto n - k - p + 1);
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1277
        val prms = map (mk_ind_perm (n - i) k p (n - l) (("a", T) :: vs))
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1278
            (j - 1 downto 0) @ prms';
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1279
        val bs = rev (List.mapPartial
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1280
          (fn ((_, T'), i) => if T = T' then SOME (Bound i) else NONE)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1281
          (List.take (vs, n - k - p - 1) ~~ (1 upto n - k - p - 1)));
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1282
        val ts = map (fn i =>
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1283
          Const ("Nominal.supp", List.nth (Ts, n - i - 1) --> S) $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1284
            fold_rev (mk_perm (T :: Ts)) prms' (Bound (n - i))) is
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1285
      in
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1286
        HOLogic.mk_Trueprop (Const ("Ex", (T --> HOLogic.boolT) --> HOLogic.boolT) $
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1287
          Abs ("a", T, HOLogic.Not $
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1288
            (Const ("op :", T --> S --> HOLogic.boolT) $ Bound 0 $
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1289
              (foldr (fn (t, u) => Const ("insert", T --> S --> S) $ t $ u)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1290
                (foldl (fn (t, u) => Const ("op Un", S --> S --> S) $ u $ t)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1291
                  (f $ Bound (n - k - p))
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1292
                  (Const ("Nominal.supp", U --> S) $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1293
                     fold_rev (mk_perm (T :: Ts)) prms (Bound (n - i')) :: ts))
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1294
                (fold_rev (mk_perm (T :: Ts)) prms (Bound (n - i - j)) :: bs)))))
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1295
      end;
18104
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1296
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1297
    fun mk_fresh_constr is p vs _ concl =
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1298
      let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1299
        val n = length vs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1300
        val Ts = map snd vs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1301
        val _ $ (_ $ _ $ t) = concl;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1302
        val c = head_of t;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1303
        val T = body_type (fastype_of c);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1304
        val k = foldr op + 0 (map (fn (_, i) => i + 1) is);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1305
        val ps = map Bound (n - k - 1 downto n - k - p);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1306
        val (_, _, ts, us) = foldl (fn ((_, i), (m, n, ts, us)) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1307
          (m - i - 1, n - i,
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1308
           ts @ map Bound (n downto n - i + 1) @ [fold_rev (mk_perm Ts)
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1309
             (map (mk_ind_perm m k p n vs) (i - 1 downto 0) @ ps)
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1310
             (Bound (m - i))],
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1311
           us @ map (fn j => fold_rev (mk_perm Ts) ps (Bound j)) (m downto m - i)))
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1312
          (n - 1, n - k - p - 2, [], []) is
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1313
      in
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1314
        HOLogic.mk_Trueprop (HOLogic.eq_const T $ list_comb (c, ts) $ list_comb (c, us))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1315
      end;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1316
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1317
    val abs_fun_finite_supp = PureThy.get_thm thy9 (Name "abs_fun_finite_supp");
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1318
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1319
    val at_finite_select = PureThy.get_thm thy9 (Name "at_finite_select");
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1320
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1321
    val induct_aux_lemmas = List.concat (map (fn Type (s, _) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1322
      [PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "_inst")),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1323
       PureThy.get_thm thy9 (Name ("fs_" ^ Sign.base_name s ^ "1")),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1324
       PureThy.get_thm thy9 (Name ("at_" ^ Sign.base_name s ^ "_inst"))]) dt_atomTs);
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1325
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1326
    val induct_aux_lemmas' = map (fn Type (s, _) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1327
      PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "2")) RS sym) dt_atomTs;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1328
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1329
    val fresh_aux = PureThy.get_thms thy9 (Name "fresh_aux");
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1330
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1331
    (**********************************************************************
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1332
      The subgoals occurring in the proof of induct_aux have the
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1333
      following parameters:
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1334
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1335
        x_1 ... x_k p_1 ... p_m z b_1 ... b_n
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1336
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1337
      where
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1338
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1339
        x_i : constructor arguments (introduced by weak induction rule)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1340
        p_i : permutations (one for each atom type in the data type)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1341
        z   : freshness context
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1342
        b_i : newly introduced names for binders (sufficiently fresh)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1343
    ***********************************************************************)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1344
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1345
    val _ = warning "proving strong induction theorem ...";
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1346
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1347
    val induct_aux = Goal.prove_global thy9 [] ind_prems' ind_concl'
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1348
      (fn prems => EVERY
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1349
        ([mk_subgoal 1 (K (K (K aux_ind_concl))),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1350
          indtac dt_induct tnames 1] @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1351
         List.concat (map (fn ((_, (_, _, constrs)), (_, constrs')) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1352
           List.concat (map (fn ((cname, cargs), is) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1353
             [simp_tac (HOL_basic_ss addsimps List.concat perm_simps') 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1354
              REPEAT (rtac allI 1)] @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1355
             List.concat (map
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1356
               (fn ((_, 0), _) => []
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1357
                 | ((i, j), k) => List.concat (map (fn j' =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1358
                     let
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1359
                       val DtType (tname, _) = List.nth (cargs, i + j');
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1360
                       val atom = Sign.base_name tname
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1361
                     in
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1362
                       [mk_subgoal 1 (mk_fresh i (i + j) j'
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1363
                          (length cargs) (length dt_atomTs)
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1364
                          (length cargs + length dt_atomTs + 1 + i - k)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1365
                          (List.mapPartial (fn (i', j) =>
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1366
                             if i = i' then NONE else SOME (i' + j)) is)),
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1367
                        rtac at_finite_select 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1368
                        rtac (PureThy.get_thm thy9 (Name ("at_" ^ atom ^ "_inst"))) 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1369
                        asm_full_simp_tac (simpset_of thy9 addsimps
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1370
                          [PureThy.get_thm thy9 (Name ("fs_" ^ atom ^ "1"))]) 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1371
                        resolve_tac prems 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1372
                        etac exE 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1373
                        asm_full_simp_tac (simpset_of thy9 addsimps
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1374
                          [symmetric fresh_def]) 1]
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1375
                     end) (0 upto j - 1))) (is ~~ (0 upto length is - 1))) @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1376
             (if exists (not o equal 0 o snd) is then
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1377
                [mk_subgoal 1 (mk_fresh_constr is (length dt_atomTs)),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1378
                 asm_full_simp_tac (simpset_of thy9 addsimps
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1379
                   (List.concat inject_thms @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1380
                    alpha @ abs_perm @ abs_fresh @ [abs_fun_finite_supp] @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1381
                    induct_aux_lemmas)) 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1382
                 dtac sym 1,
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1383
                 asm_full_simp_tac (simpset_of thy9) 1,
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1384
                 REPEAT (etac conjE 1)]
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1385
              else
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1386
                []) @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1387
             [(resolve_tac prems THEN_ALL_NEW
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1388
                (atac ORELSE'
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1389
                  SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1390
                      _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1391
                        asm_simp_tac (simpset_of thy9 addsimps fresh_aux) i
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1392
                    | _ =>
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1393
                        asm_simp_tac (simpset_of thy9
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1394
                        addsimps induct_aux_lemmas'
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1395
                        addsimprocs [perm_simproc]) i))) 1])
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1396
               (constrs ~~ constrs'))) (descr'' ~~ ndescr)) @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1397
         [REPEAT (eresolve_tac [conjE, allE_Nil] 1),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1398
          REPEAT (etac allE 1),
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1399
          REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac (simpset_of thy9) 1)]));
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1400
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1401
    val induct_aux' = Thm.instantiate ([],
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1402
      map (fn (s, T) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1403
        let val pT = TVar (("'n", 0), HOLogic.typeS) --> T --> HOLogic.boolT
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1404
        in (cterm_of thy9 (Var ((s, 0), pT)), cterm_of thy9 (Free (s, pT))) end)
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1405
          (pnames ~~ recTs) @
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1406
      map (fn (_, f) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1407
        let val f' = Logic.varify f
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1408
        in (cterm_of thy9 f',
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
  1409
          cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1410
        end) fresh_fs) induct_aux;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1411
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1412
    val induct = Goal.prove_global thy9 [] ind_prems ind_concl
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1413
      (fn prems => EVERY
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1414
         [rtac induct_aux' 1,
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1415
          REPEAT (resolve_tac induct_aux_lemmas 1),
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1416
          REPEAT ((resolve_tac prems THEN_ALL_NEW
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1417
            (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1418
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1419
    val (_, thy10) = thy9 |>
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1420
      Theory.add_path big_name |>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1421
      PureThy.add_thms [(("induct'", induct_aux), [])] ||>>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1422
      PureThy.add_thms [(("induct", induct), [case_names_induct])] ||>>
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1423
      PureThy.add_thmss [(("inducts", projections induct), [case_names_induct])];
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1424
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1425
    (**** recursion combinator ****)
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1426
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1427
    val _ = warning "defining recursion combinator ...";
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1428
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1429
    val used = foldr add_typ_tfree_names [] recTs;
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1430
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1431
    val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts' used;
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1432
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  1433
    val rec_sort = if null dt_atomTs then HOLogic.typeS else
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1434
      let val names = map (Sign.base_name o fst o dest_Type) dt_atomTs
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1435
      in Sign.certify_sort thy10 (map (Sign.intern_class thy10)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1436
        (map (fn s => "pt_" ^ s) names @
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1437
         List.concat (map (fn s => List.mapPartial (fn s' =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1438
           if s = s' then NONE
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1439
           else SOME ("cp_" ^ s ^ "_" ^ s')) names) names)))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1440
      end;
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1441
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1442
    val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1443
    val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1444
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1445
    val rec_set_Ts = map (fn (T1, T2) =>
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1446
      rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1447
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1448
    val big_rec_name = big_name ^ "_rec_set";
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1449
    val rec_set_names' =
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1450
      if length descr'' = 1 then [big_rec_name] else
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1451
        map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1452
          (1 upto (length descr''));
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1453
    val rec_set_names =  map (Sign.full_name (Theory.sign_of thy10)) rec_set_names';
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1454
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1455
    val rec_fns = map (uncurry (mk_Free "f"))
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1456
      (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1457
    val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1458
      (rec_set_names' ~~ rec_set_Ts);
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1459
    val rec_sets = map (fn c => list_comb (Const c, rec_fns))
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1460
      (rec_set_names ~~ rec_set_Ts);
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1461
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1462
    (* introduction rules for graph of recursion function *)
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1463
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1464
    val rec_preds = map (fn (a, T) =>
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1465
      Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1466
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1467
    fun mk_fresh3 rs [] = []
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1468
      | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) =>
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1469
            List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1470
              else SOME (HOLogic.mk_Trueprop
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1471
                (Const ("Nominal.fresh", T --> U --> HOLogic.boolT) $ Free y $ Free r)))
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1472
                  rs) ys) @
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1473
          mk_fresh3 rs yss;
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1474
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1475
    (* FIXME: avoid collisions with other variable names? *)
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1476
    val rec_ctxt = Free ("z", fsT');
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1477
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1478
    fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1479
          rec_eq_prems, l), ((cname, cargs), idxs)) =
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1480
      let
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1481
        val Ts = map (typ_of_dtyp descr'' sorts') cargs;
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1482
        val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1483
        val frees' = partition_cargs idxs frees;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1484
        val binders = List.concat (map fst frees');
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1485
        val atomTs = distinct op = (maps (map snd o fst) frees');
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1486
        val recs = List.mapPartial
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1487
          (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE)
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1488
          (partition_cargs idxs cargs ~~ frees');
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1489
        val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1490
          map (fn (i, _) => List.nth (rec_result_Ts, i)) recs;
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1491
        val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1492
          (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees'');
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1493
        val prems2 =
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1494
          map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1495
            (Const ("Nominal.fresh", T --> fastype_of f --> HOLogic.boolT) $
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1496
              Free p $ f)) binders) rec_fns;
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1497
        val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1498
        val prems4 = map (fn ((i, _), y) =>
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1499
          HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees'');
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1500
        val prems5 = mk_fresh3 (recs ~~ frees'') frees';
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1501
        val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1502
          (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1503
             (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y)))
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1504
               frees'') atomTs;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1505
        val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1506
          (Const ("Nominal.fresh", T --> fsT' --> HOLogic.boolT) $
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1507
             Free x $ rec_ctxt)) binders;
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1508
        val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees''));
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1509
        val result_freshs = map (fn p as (_, T) =>
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1510
          Const ("Nominal.fresh", T --> fastype_of result --> HOLogic.boolT) $
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1511
            Free p $ result) binders;
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1512
        val P = HOLogic.mk_Trueprop (p $ result)
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1513
      in
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1514
        (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1515
           HOLogic.mk_Trueprop (rec_set $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1516
             list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1517
         rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))],
21054
6048c085c3ae Split up FCBs into separate formulae for each binder.
berghofe
parents: 21021
diff changeset
  1518
         rec_prems' @ map (fn fr => list_all_free (frees @ frees'',
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1519
           Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
21054
6048c085c3ae Split up FCBs into separate formulae for each binder.
berghofe
parents: 21021
diff changeset
  1520
             HOLogic.mk_Trueprop fr))) result_freshs,
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1521
         rec_eq_prems @ [List.concat prems2 @ prems3],
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1522
         l + 1)
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1523
      end;
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1524
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1525
    val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1526
      Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1527
        Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1528
          (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1529
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  1530
    val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
21055
e053811d680e Induction rule for graph of recursion combinator
berghofe
parents: 21054
diff changeset
  1531
      thy10 |>
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1532
      setmp InductivePackage.quiet_mode (!quiet_mode)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1533
        (TheoryTarget.init NONE #>
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1534
         InductivePackage.add_inductive_i false big_rec_name false false false
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1535
           (map (fn (s, T) => (s, SOME T, NoSyn)) (rec_set_names' ~~ rec_set_Ts))
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1536
           (map (apsnd SOME o dest_Free) rec_fns)
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1537
           (map (fn x => (("", []), x)) rec_intr_ts) [] #>
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  1538
         apsnd (ProofContext.theory_of o LocalTheory.exit)) ||>
21055
e053811d680e Induction rule for graph of recursion combinator
berghofe
parents: 21054
diff changeset
  1539
      PureThy.hide_thms true [NameSpace.append
e053811d680e Induction rule for graph of recursion combinator
berghofe
parents: 21054
diff changeset
  1540
        (Sign.full_name thy10 big_rec_name) "induct"];
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1541
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1542
    (** equivariance **)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1543
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1544
    val fresh_bij = PureThy.get_thms thy11 (Name "fresh_bij");
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1545
    val perm_bij = PureThy.get_thms thy11 (Name "perm_bij");
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1546
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1547
    val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1548
      let
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1549
        val permT = mk_permT aT;
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1550
        val pi = Free ("pi", permT);
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1551
        val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f"))
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1552
          (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1553
        val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1554
          (rec_set_names ~~ rec_set_Ts);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1555
        val ps = map (fn ((((T, U), R), R'), i) =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1556
          let
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1557
            val x = Free ("x" ^ string_of_int i, T);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1558
            val y = Free ("y" ^ string_of_int i, U)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1559
          in
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1560
            (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1561
          end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1562
        val ths = map (fn th => standard (th RS mp)) (split_conj_thm
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1563
          (Goal.prove_global thy11 [] []
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1564
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps)))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1565
            (fn _ => rtac rec_induct 1 THEN REPEAT
22433
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
  1566
               (NominalPermeq.perm_simp_tac (HOL_basic_ss addsimps flat perm_simps') 1 THEN
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1567
                (resolve_tac rec_intrs THEN_ALL_NEW
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1568
                 asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1569
        val ths' = map (fn ((P, Q), th) =>
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1570
          Goal.prove_global thy11 [] []
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1571
            (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1572
            (fn _ => dtac (Thm.instantiate ([],
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1573
                 [(cterm_of thy11 (Var (("pi", 0), permT)),
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1574
                   cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1575
               NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1576
      in (ths, ths') end) dt_atomTs);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1577
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1578
    (** finite support **)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1579
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1580
    val rec_fin_supp_thms = map (fn aT =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1581
      let
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1582
        val name = Sign.base_name (fst (dest_Type aT));
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1583
        val fs_name = PureThy.get_thm thy11 (Name ("fs_" ^ name ^ "1"));
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1584
        val aset = HOLogic.mk_setT aT;
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1585
        val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1586
        val fins = map (fn (f, T) => HOLogic.mk_Trueprop
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1587
          (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1588
            (rec_fns ~~ rec_fn_Ts)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1589
      in
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1590
        map (fn th => standard (th RS mp)) (split_conj_thm
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1591
          (Goal.prove_global thy11 [] fins
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1592
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1593
              (map (fn (((T, U), R), i) =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1594
                 let
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1595
                   val x = Free ("x" ^ string_of_int i, T);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1596
                   val y = Free ("y" ^ string_of_int i, U)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1597
                 in
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1598
                   HOLogic.mk_imp (R $ x $ y,
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1599
                     finite $ (Const ("Nominal.supp", U --> aset) $ y))
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1600
                 end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ (1 upto length recTs)))))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1601
            (fn fins => (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1602
               (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1603
      end) dt_atomTs;
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1604
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1605
    (** freshness **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1606
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1607
    val perm_fresh_fresh = PureThy.get_thms thy11 (Name "perm_fresh_fresh");
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1608
    val perm_swap = PureThy.get_thms thy11 (Name "perm_swap");
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1609
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1610
    val finite_premss = map (fn aT =>
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1611
      map (fn (f, T) => HOLogic.mk_Trueprop
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1612
        (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1613
           (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1614
           (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1615
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1616
    val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1617
      let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1618
        val name = Sign.base_name (fst (dest_Type aT));
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1619
        val fs_name = PureThy.get_thm thy11 (Name ("fs_" ^ name ^ "1"));
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1620
        val a = Free ("a", aT);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1621
        val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1622
            (Const ("Nominal.fresh", aT --> fT --> HOLogic.boolT) $ a $ f))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1623
          (rec_fns ~~ rec_fn_Ts)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1624
      in
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1625
        map (fn (((T, U), R), eqvt_th) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1626
          let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1627
            val x = Free ("x", T);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1628
            val y = Free ("y", U);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1629
            val y' = Free ("y'", U)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1630
          in
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21377
diff changeset
  1631
            standard (Goal.prove (ProofContext.init thy11) [] (finite_prems @
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1632
                [HOLogic.mk_Trueprop (R $ x $ y),
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1633
                 HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1634
                   HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1635
                 HOLogic.mk_Trueprop (Const ("Nominal.fresh",
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1636
                   aT --> T --> HOLogic.boolT) $ a $ x)] @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1637
              freshs)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1638
              (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1639
                 aT --> U --> HOLogic.boolT) $ a $ y))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1640
              (fn {prems, context} =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1641
                 let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1642
                   val (finite_prems, rec_prem :: unique_prem ::
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1643
                     fresh_prems) = chop (length finite_prems) prems;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1644
                   val unique_prem' = unique_prem RS spec RS mp;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1645
                   val unique = [unique_prem', unique_prem' RS sym] MRS trans;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1646
                   val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1647
                   val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1648
                 in EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1649
                   [rtac (Drule.cterm_instantiate
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1650
                      [(cterm_of thy11 S,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1651
                        cterm_of thy11 (Const ("Nominal.supp",
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1652
                          fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1653
                      supports_fresh) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1654
                    simp_tac (HOL_basic_ss addsimps
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1655
                      [supports_def, symmetric fresh_def, fresh_prod]) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1656
                    REPEAT_DETERM (resolve_tac [allI, impI] 1),
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1657
                    REPEAT_DETERM (etac conjE 1),
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1658
                    rtac unique 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1659
                    SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1660
                      [cut_facts_tac [rec_prem] 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1661
                       rtac (Thm.instantiate ([],
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1662
                         [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1663
                           cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1664
                       asm_simp_tac (HOL_ss addsimps
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1665
                         (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1666
                    rtac rec_prem 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1667
                    simp_tac (HOL_ss addsimps (fs_name ::
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1668
                      supp_prod :: finite_Un :: finite_prems)) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1669
                    simp_tac (HOL_ss addsimps (symmetric fresh_def ::
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1670
                      fresh_prod :: fresh_prems)) 1]
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1671
                 end))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1672
          end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1673
      end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1674
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1675
    (** uniqueness **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1676
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21365
diff changeset
  1677
    val exists_fresh' = PureThy.get_thms thy11 (Name "exists_fresh'");
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1678
    val fs_atoms = map (fn Type (s, _) => PureThy.get_thm thy11
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1679
      (Name ("fs_" ^ Sign.base_name s ^ "1"))) dt_atomTs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1680
    val fresh_atm = PureThy.get_thms thy11 (Name "fresh_atm");
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1681
    val calc_atm = PureThy.get_thms thy11 (Name "calc_atm");
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1682
    val fresh_left = PureThy.get_thms thy11 (Name "fresh_left");
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1683
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1684
    val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1685
    val fun_tupleT = fastype_of fun_tuple;
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1686
    val rec_unique_frees =
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1687
      DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs;
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1688
    val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1689
    val rec_unique_frees' =
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1690
      DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1691
    val rec_unique_concls = map (fn ((x, U), R) =>
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1692
        Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1693
          Abs ("y", U, R $ Free x $ Bound 0))
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1694
      (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1695
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1696
    val induct_aux_rec = Drule.cterm_instantiate
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1697
      (map (pairself (cterm_of thy11))
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1698
         (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1699
            Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1700
              fresh_fs @
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1701
          map (fn (((P, T), (x, U)), Q) =>
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1702
           (Var ((P, 0), HOLogic.unitT --> Logic.varifyT T --> HOLogic.boolT),
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1703
            Abs ("z", HOLogic.unitT, absfree (x, U, Q))))
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1704
              (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1705
          map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T)))
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1706
            rec_unique_frees)) induct_aux;
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1707
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1708
    fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1709
      let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1710
        val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1711
        val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1712
            (HOLogic.exists_const T $ Abs ("x", T,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1713
              Const ("Nominal.fresh", T --> fastype_of p --> HOLogic.boolT) $
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1714
                Bound 0 $ p)))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1715
          (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1716
            [cut_facts_tac ths 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1717
             REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21365
diff changeset
  1718
             resolve_tac exists_fresh' 1,
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1719
             asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1720
        val (([cx], ths), ctxt') = Obtain.result
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1721
          (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1722
            [etac exE 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1723
             full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1724
             REPEAT (etac conjE 1)])
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1725
          [ex] ctxt
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1726
      in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1727
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1728
    val finite_ctxt_prems = map (fn aT =>
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1729
      HOLogic.mk_Trueprop
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1730
        (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1731
           (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1732
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1733
    val rec_unique_thms = split_conj_thm (Goal.prove
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 21377
diff changeset
  1734
      (ProofContext.init thy11) (map fst rec_unique_frees)
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1735
      (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems')
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1736
      (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1737
      (fn {prems, context} =>
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1738
         let
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1739
           val k = length rec_fns;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1740
           val (finite_thss, ths1) = fold_map (fn T => fn xs =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1741
             apfst (pair T) (chop k xs)) dt_atomTs prems;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1742
           val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1743
           val (P_ind_ths, fcbs) = chop k ths2;
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1744
           val P_ths = map (fn th => th RS mp) (split_conj_thm
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1745
             (Goal.prove context
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1746
               (map fst (rec_unique_frees'' @ rec_unique_frees')) []
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1747
               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1748
                  (map (fn (((x, y), S), P) => HOLogic.mk_imp
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1749
                    (S $ Free x $ Free y, P $ (Free y)))
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1750
                      (rec_unique_frees'' ~~ rec_unique_frees' ~~ rec_sets ~~ rec_preds))))
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1751
               (fn _ =>
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1752
                  rtac rec_induct 1 THEN
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1753
                  REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1754
           val rec_fin_supp_thms' = map
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1755
             (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1756
             (rec_fin_supp_thms ~~ finite_thss);
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1757
         in EVERY
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1758
           ([rtac induct_aux_rec 1] @
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1759
            maps (fn ((_, finite_ths), finite_th) =>
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1760
              [cut_facts_tac (finite_th :: finite_ths) 1,
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1761
               asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1762
                (finite_thss ~~ finite_ctxt_ths) @
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1763
            maps (fn ((_, idxss), elim) => maps (fn idxs =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1764
              [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1765
               REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1766
               rtac ex1I 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1767
               (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1768
               rotate_tac ~1 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1769
               ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1770
                  (HOL_ss addsimps List.concat distinct_thms)) 1] @
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1771
               (if null idxs then [] else [hyp_subst_tac 1,
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1772
                SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1773
                  let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1774
                    val SOME prem = find_first (can (HOLogic.dest_eq o
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1775
                      HOLogic.dest_Trueprop o prop_of)) prems';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1776
                    val _ $ (_ $ lhs $ rhs) = prop_of prem;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1777
                    val _ $ (_ $ lhs' $ rhs') = term_of concl;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1778
                    val rT = fastype_of lhs';
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1779
                    val (c, cargsl) = strip_comb lhs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1780
                    val cargsl' = partition_cargs idxs cargsl;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1781
                    val boundsl = List.concat (map fst cargsl');
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1782
                    val (_, cargsr) = strip_comb rhs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1783
                    val cargsr' = partition_cargs idxs cargsr;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1784
                    val boundsr = List.concat (map fst cargsr');
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1785
                    val (params1, _ :: params2) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1786
                      chop (length params div 2) (map term_of params);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1787
                    val params' = params1 @ params2;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1788
                    val rec_prems = filter (fn th => case prop_of th of
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1789
                      _ $ (S $ _ $ _) => S mem rec_sets | _ => false) prems';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1790
                    val fresh_prems = filter (fn th => case prop_of th of
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1791
                        _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1792
                      | _ $ (Const ("Not", _) $ _) => true
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1793
                      | _ => false) prems';
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1794
                    val Ts = map fastype_of boundsl;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1795
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1796
                    val _ = warning "step 1: obtaining fresh names";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1797
                    val (freshs1, freshs2, context'') = fold
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1798
                      (obtain_fresh_name (rec_ctxt :: rec_fns @ params')
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1799
                         (List.concat (map snd finite_thss) @
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1800
                            finite_ctxt_ths @ rec_prems)
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1801
                         rec_fin_supp_thms')
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1802
                      Ts ([], [], context');
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1803
                    val pi1 = map perm_of_pair (boundsl ~~ freshs1);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1804
                    val rpi1 = rev pi1;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1805
                    val pi2 = map perm_of_pair (boundsr ~~ freshs1);
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1806
                    val rpi2 = rev pi2;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1807
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1808
                    val fresh_prems' = mk_not_sym fresh_prems;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1809
                    val freshs2' = mk_not_sym freshs2;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1810
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1811
                    (** as, bs, cs # K as ts, K bs us **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1812
                    val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1813
                    val prove_fresh_ss = HOL_ss addsimps
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1814
                      (finite_Diff :: List.concat fresh_thms @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1815
                       fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1816
                    (* FIXME: avoid asm_full_simp_tac ? *)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1817
                    fun prove_fresh ths y x = Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1818
                      (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1819
                         fastype_of x --> fastype_of y --> HOLogic.boolT) $ x $ y))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1820
                      (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1821
                    val constr_fresh_thms =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1822
                      map (prove_fresh fresh_prems lhs) boundsl @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1823
                      map (prove_fresh fresh_prems rhs) boundsr @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1824
                      map (prove_fresh freshs2 lhs) freshs1 @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1825
                      map (prove_fresh freshs2 rhs) freshs1;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1826
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1827
                    (** pi1 o (K as ts) = pi2 o (K bs us) **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1828
                    val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1829
                    val pi1_pi2_eq = Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1830
                      (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1831
                        (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1832
                      (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1833
                         [cut_facts_tac constr_fresh_thms 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1834
                          asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1835
                          rtac prem 1]);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1836
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1837
                    (** pi1 o ts = pi2 o us **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1838
                    val _ = warning "step 4: pi1 o ts = pi2 o us";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1839
                    val pi1_pi2_eqs = map (fn (t, u) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1840
                      Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1841
                        (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1842
                          (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1843
                        (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1844
                           [cut_facts_tac [pi1_pi2_eq] 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1845
                            asm_full_simp_tac (HOL_ss addsimps
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1846
                              (calc_atm @ List.concat perm_simps' @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1847
                               fresh_prems' @ freshs2' @ abs_perm @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1848
                               alpha @ List.concat inject_thms)) 1]))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1849
                        (map snd cargsl' ~~ map snd cargsr');
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1850
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1851
                    (** pi1^-1 o pi2 o us = ts **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1852
                    val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1853
                    val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1854
                      Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1855
                        (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1856
                          (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1857
                        (fn _ => simp_tac (HOL_ss addsimps
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1858
                           ((eq RS sym) :: perm_swap)) 1))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1859
                        (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1860
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1861
                    val (rec_prems1, rec_prems2) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1862
                      chop (length rec_prems div 2) rec_prems;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1863
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1864
                    (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1865
                    val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1866
                    val rec_prems' = map (fn th =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1867
                      let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1868
                        val _ $ (S $ x $ y) = prop_of th;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1869
                        val k = find_index (equal S) rec_sets;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1870
                        val pi = rpi1 @ pi2;
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1871
                        fun mk_pi z = fold_rev (mk_perm []) pi z;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1872
                        fun eqvt_tac p =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1873
                          let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1874
                            val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1875
                            val l = find_index (equal T) dt_atomTs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1876
                            val th = List.nth (List.nth (rec_equiv_thms', l), k);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1877
                            val th' = Thm.instantiate ([],
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1878
                              [(cterm_of thy11 (Var (("pi", 0), U)),
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1879
                                cterm_of thy11 p)]) th;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1880
                          in rtac th' 1 end;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1881
                        val th' = Goal.prove context'' [] []
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1882
                          (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1883
                          (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1884
                             (map eqvt_tac pi @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1885
                              [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1886
                                 perm_swap @ perm_fresh_fresh)) 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1887
                               rtac th 1]))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1888
                      in
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1889
                        Simplifier.simplify
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1890
                          (HOL_basic_ss addsimps rpi1_pi2_eqs) th'
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1891
                      end) rec_prems2;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1892
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1893
                    val ihs = filter (fn th => case prop_of th of
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1894
                      _ $ (Const ("All", _) $ _) => true | _ => false) prems';
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1895
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1896
                    (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1897
                    val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1898
                    val rec_eqns = map (fn (th, ih) =>
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1899
                      let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1900
                        val th' = th RS (ih RS spec RS mp) RS sym;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1901
                        val _ $ (_ $ lhs $ rhs) = prop_of th';
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1902
                        fun strip_perm (_ $ _ $ t) = strip_perm t
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1903
                          | strip_perm t = t;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1904
                      in
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1905
                        Goal.prove context'' [] []
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1906
                           (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1907
                              (fold_rev (mk_perm []) pi1 lhs,
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1908
                               fold_rev (mk_perm []) pi2 (strip_perm rhs))))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1909
                           (fn _ => simp_tac (HOL_basic_ss addsimps
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1910
                              (th' :: perm_swap)) 1)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1911
                      end) (rec_prems' ~~ ihs);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1912
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1913
                    (** as # rs **)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1914
                    val _ = warning "step 8: as # rs";
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1915
                    val rec_freshs = List.concat
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1916
                      (map (fn (rec_prem, ih) =>
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1917
                        let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1918
                          val _ $ (S $ x $ (y as Free (_, T))) =
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1919
                            prop_of rec_prem;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1920
                          val k = find_index (equal S) rec_sets;
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1921
                          val atoms = List.concat (List.mapPartial (fn (bs, z) =>
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1922
                            if z = x then NONE else SOME bs) cargsl')
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1923
                        in
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1924
                          map (fn a as Free (_, aT) =>
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1925
                            let val l = find_index (equal aT) dt_atomTs;
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1926
                            in
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1927
                              Goal.prove context'' [] []
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1928
                                (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1929
                                  aT --> T --> HOLogic.boolT) $ a $ y))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1930
                                (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1931
                                   (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 ::
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1932
                                    map (fn th => rtac th 1)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1933
                                      (snd (List.nth (finite_thss, l))) @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1934
                                    [rtac rec_prem 1, rtac ih 1,
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1935
                                     REPEAT_DETERM (resolve_tac fresh_prems 1)]))
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1936
                            end) atoms
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1937
                        end) (rec_prems1 ~~ ihs));
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1938
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1939
                    (** as # fK as ts rs , bs # fK bs us vs **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1940
                    val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1941
                    fun prove_fresh_result (a as Free (_, aT)) =
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1942
                      Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1943
                        (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1944
                          aT --> rT --> HOLogic.boolT) $ a $ rhs'))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1945
                        (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1946
                           [resolve_tac fcbs 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1947
                            REPEAT_DETERM (resolve_tac
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1948
                              (fresh_prems @ rec_freshs) 1),
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1949
                            REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1950
                              THEN resolve_tac rec_prems 1),
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1951
                            resolve_tac P_ind_ths 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1952
                            REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  1953
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1954
                    val fresh_results'' = map prove_fresh_result boundsl;
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1955
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1956
                    fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1957
                      let val th' = Goal.prove context'' [] []
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1958
                        (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1959
                          aT --> rT --> HOLogic.boolT) $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1960
                            fold_rev (mk_perm []) (rpi2 @ pi1) a $
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1961
                            fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1962
                        (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1963
                           rtac th 1)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1964
                      in
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1965
                        Goal.prove context'' [] []
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1966
                          (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1967
                            aT --> rT --> HOLogic.boolT) $ b $ lhs'))
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1968
                          (fn _ => EVERY
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1969
                             [cut_facts_tac [th'] 1,
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1970
                              NominalPermeq.perm_simp_tac (HOL_ss addsimps
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1971
                                (rec_eqns @ pi1_pi2_eqs @ perm_swap)) 1,
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1972
                              full_simp_tac (HOL_ss addsimps (calc_atm @
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1973
                                fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1974
                      end;
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1975
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1976
                    val fresh_results = fresh_results'' @ map prove_fresh_result''
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1977
                      (boundsl ~~ boundsr ~~ fresh_results'');
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1978
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1979
                    (** cs # fK as ts rs , cs # fK bs us vs **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1980
                    val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1981
                    fun prove_fresh_result' recs t (a as Free (_, aT)) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1982
                      Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1983
                        (HOLogic.mk_Trueprop (Const ("Nominal.fresh",
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1984
                          aT --> rT --> HOLogic.boolT) $ a $ t))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1985
                        (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1986
                          [cut_facts_tac recs 1,
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1987
                           REPEAT_DETERM (dresolve_tac
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1988
                             (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1989
                           NominalPermeq.fresh_guess_tac
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1990
                             (HOL_ss addsimps (freshs2 @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1991
                                fs_atoms @ fresh_atm @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1992
                                List.concat (map snd finite_thss))) 1]);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1993
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1994
                    val fresh_results' =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1995
                      map (prove_fresh_result' rec_prems1 rhs') freshs1 @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1996
                      map (prove_fresh_result' rec_prems2 lhs') freshs1;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1997
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1998
                    (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1999
                    val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2000
                    val pi1_pi2_result = Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2001
                      (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  2002
                        (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2003
                      (fn _ => NominalPermeq.perm_simp_tac (HOL_ss addsimps
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  2004
                           pi1_pi2_eqs @ rec_eqns) 1 THEN
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2005
                         TRY (simp_tac (HOL_ss addsimps
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2006
                           (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2007
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2008
                    val _ = warning "final result";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2009
                    val final = Goal.prove context'' [] [] (term_of concl)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2010
                      (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2011
                        full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2012
                          fresh_results @ fresh_results') 1);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2013
                    val final' = ProofContext.export context'' context' [final];
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2014
                    val _ = warning "finished!"
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2015
                  in
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2016
                    resolve_tac final' 1
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2017
                  end) context 1])) idxss) (ndescr ~~ rec_elims))
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2018
         end));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2019
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2020
    val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2021
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2022
    (* define primrec combinators *)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2023
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2024
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2025
    val reccomb_names = map (Sign.full_name thy11)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2026
      (if length descr'' = 1 then [big_reccomb_name] else
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2027
        (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2028
          (1 upto (length descr''))));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2029
    val reccombs = map (fn ((name, T), T') => list_comb
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2030
      (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2031
        (reccomb_names ~~ recTs ~~ rec_result_Ts);
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2032
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2033
    val (reccomb_defs, thy12) =
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2034
      thy11
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2035
      |> Theory.add_consts_i (map (fn ((name, T), T') =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2036
          (Sign.base_name name, rec_fn_Ts @ [T] ---> T', NoSyn))
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2037
          (reccomb_names ~~ recTs ~~ rec_result_Ts))
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2038
      |> (PureThy.add_defs_i false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2039
          ((Sign.base_name name) ^ "_def", Logic.mk_equals (comb, absfree ("x", T,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2040
           Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  2041
             set $ Free ("x", T) $ Free ("y", T'))))))
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2042
               (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2043
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2044
    (* prove characteristic equations for primrec combinators *)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2045
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2046
    val rec_thms = map (fn (prems, concl) =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2047
      let
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2048
        val _ $ (_ $ (_ $ x) $ _) = concl;
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2049
        val (_, cargs) = strip_comb x;
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2050
        val ps = map (fn (x as Free (_, T), i) =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2051
          (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2052
        val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  2053
        val prems' = List.concat finite_premss @ finite_ctxt_prems @
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  2054
          rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2055
        fun solve rules prems = resolve_tac rules THEN_ALL_NEW
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2056
          (resolve_tac prems THEN_ALL_NEW atac)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2057
      in
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2058
        Goal.prove_global thy12 [] prems' concl'
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2059
          (fn prems => EVERY
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2060
            [rewrite_goals_tac reccomb_defs,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2061
             rtac the1_equality 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2062
             solve rec_unique_thms prems 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2063
             resolve_tac rec_intrs 1,
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2064
             REPEAT (solve (prems @ rec_total_thms) prems 1)])
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2065
      end) (rec_eq_prems ~~
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2066
        DatatypeProp.make_primrecs new_type_names descr' sorts' thy12);
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  2067
22433
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
  2068
    val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
  2069
      ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
  2070
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  2071
    (* FIXME: theorems are stored in database for testing only *)
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2072
    val (_, thy13) = thy12 |>
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  2073
      PureThy.add_thmss
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  2074
        [(("rec_equiv", List.concat rec_equiv_thms), []),
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  2075
         (("rec_equiv'", List.concat rec_equiv_thms'), []),
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  2076
         (("rec_fin_supp", List.concat rec_fin_supp_thms), []),
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2077
         (("rec_fresh", List.concat rec_fresh_thms), []),
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2078
         (("rec_unique", map standard rec_unique_thms), []),
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2079
         (("recs", rec_thms), [])] ||>
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
  2080
      Theory.parent_path ||>
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
  2081
      map_nominal_datatypes (fold Symtab.update dt_infos);
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  2082
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2083
  in
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2084
    thy13
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2085
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2086
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2087
val add_nominal_datatype = gen_add_nominal_datatype read_typ true;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2088
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2089
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2090
(* FIXME: The following stuff should be exported by DatatypePackage *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2091
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2092
local structure P = OuterParse and K = OuterKeyword in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2093
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2094
val datatype_decl =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2095
  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2096
    (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2097
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2098
fun mk_datatype args =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2099
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2100
    val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2101
    val specs = map (fn ((((_, vs), t), mx), cons) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2102
      (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
  2103
  in add_nominal_datatype false names specs end;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2104
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2105
val nominal_datatypeP =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2106
  OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2107
    (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2108
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2109
val _ = OuterSyntax.add_parsers [nominal_datatypeP];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2110
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2111
end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2112
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
  2113
end