src/HOL/Nominal/nominal_package.ML
author urbanc
Fri, 25 Nov 2005 14:51:39 +0100
changeset 18246 676d2e625d98
parent 18245 65e60434b3c2
child 18261 1318955d57ac
permissions -rw-r--r--
added fsub.thy (poplmark challenge) to the examples directory.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     1
(* $Id$ *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     3
signature NOMINAL_PACKAGE =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     4
sig
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     5
  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
     6
    (bstring * string list * mixfix) list) list -> theory -> theory
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     7
end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     8
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
     9
structure NominalPackage : NOMINAL_PACKAGE =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    10
struct
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    11
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    12
open DatatypeAux;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
    13
open NominalAtoms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    14
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    15
(** FIXME: DatatypePackage should export this function **)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    16
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    17
local
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    18
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    19
fun dt_recs (DtTFree _) = []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    20
  | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    21
  | dt_recs (DtRec i) = [i];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    22
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    23
fun dt_cases (descr: descr) (_, args, constrs) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    24
  let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    25
    fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    26
    val bnames = map the_bname (distinct (List.concat (map dt_recs args)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    27
  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
    28
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    29
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    30
fun induct_cases descr =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    31
  DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    32
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    33
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
    34
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    35
in
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    36
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    37
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    38
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    39
fun mk_case_names_exhausts descr new =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    40
  map (RuleCases.case_names o exhaust_cases descr o #1)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    41
    (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
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
end;
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
(*******************************)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    46
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    47
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    48
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    49
fun read_typ sign ((Ts, sorts), str) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    50
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    51
    val T = Type.no_tvars (Sign.read_typ (sign, (AList.lookup op =)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    52
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    53
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    54
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    55
(** taken from HOL/Tools/datatype_aux.ML **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    56
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    57
fun indtac indrule indnames i st =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    58
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    59
    val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    60
    val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    61
      (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    62
    val getP = if can HOLogic.dest_imp (hd ts) then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    63
      (apfst SOME) o HOLogic.dest_imp else pair NONE;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    64
    fun abstr (t1, t2) = (case t1 of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    65
        NONE => (case filter (fn Free (s, _) => s mem indnames | _ => false)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    66
              (term_frees t2) of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    67
            [Free (s, T)] => absfree (s, T, t2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    68
          | _ => sys_error "indtac")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    69
      | SOME (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    70
    val cert = cterm_of (Thm.sign_of_thm st);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    71
    val Ps = map (cert o head_of o snd o getP) ts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    72
    val indrule' = cterm_instantiate (Ps ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    73
      (map (cert o abstr o getP) ts')) indrule
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    74
  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    75
    rtac indrule' i st
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    76
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    77
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    78
fun gen_add_nominal_datatype prep_typ err flat_names new_type_names dts thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    79
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    80
    (* this theory is used just for parsing *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    81
  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    82
    val tmp_thy = thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    83
      Theory.copy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    84
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    85
        (tname, length tvs, mx)) dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    86
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    87
    val sign = Theory.sign_of tmp_thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    88
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    89
    val atoms = atoms_of thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    90
    val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    91
    val cp_classes = List.concat (map (fn atom1 => map (fn atom2 =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    92
      Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    93
        Sign.base_name atom2)) atoms) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    94
    fun augment_sort S = S union classes;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    95
    val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    96
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    97
    fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    98
      let val (cargs', sorts') = Library.foldl (prep_typ sign) (([], sorts), cargs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    99
      in (constrs @ [(cname, cargs', mx)], sorts') end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   100
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   101
    fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   102
      let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   103
      in (dts @ [(tvs, tname, mx, constrs')], sorts') end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   104
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   105
    val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   106
    val sorts' = map (apsnd augment_sort) sorts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   107
    val tyvars = map #1 dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   108
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   109
    val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   110
    val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   111
      map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   112
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   113
    val ps = map (fn (_, n, _, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   114
      (Sign.full_name sign n, Sign.full_name sign (n ^ "_Rep"))) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   115
    val rps = map Library.swap ps;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   116
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   117
    fun replace_types (Type ("nominal.ABS", [T, U])) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   118
          Type ("fun", [T, Type ("nominal.nOption", [replace_types U])])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   119
      | replace_types (Type (s, Ts)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   120
          Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   121
      | replace_types T = T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   122
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   123
    fun replace_types' (Type (s, Ts)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   124
          Type (getOpt (AList.lookup op = rps s, s), map replace_types' Ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   125
      | replace_types' T = T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   126
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   127
    val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   128
      map (fn (cname, cargs, mx) => (cname,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   129
        map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   130
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   131
    val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   132
    val full_new_type_names' = map (Sign.full_name (sign_of thy)) new_type_names';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   133
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   134
    val ({induction, ...},thy1) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   135
      DatatypePackage.add_datatype_i err flat_names new_type_names' dts'' thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   136
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   137
    val SOME {descr, ...} = Symtab.lookup
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   138
      (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   139
    fun nth_dtyp i = typ_of_dtyp descr sorts' (DtRec i);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   140
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   141
    val dt_atomTs = filter (fn Type (s, []) => s mem atoms | _ => false)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   142
      (get_nonrec_types descr sorts);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   143
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   144
    (**** define permutation functions ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   145
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   146
    val permT = mk_permT (TFree ("'x", HOLogic.typeS));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   147
    val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   148
    val perm_types = map (fn (i, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   149
      let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   150
      in permT --> T --> T end) descr;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   151
    val perm_names = replicate (length new_type_names) "nominal.perm" @
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   152
      DatatypeProp.indexify_names (map (fn i => Sign.full_name (sign_of thy1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   153
        ("perm_" ^ name_of_typ (nth_dtyp i)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   154
          (length new_type_names upto length descr - 1));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   155
    val perm_names_types = perm_names ~~ perm_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   156
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   157
    val perm_eqs = List.concat (map (fn (i, (_, _, constrs)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   158
      let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   159
      in map (fn (cname, dts) => 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   160
        let
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   161
          val Ts = map (typ_of_dtyp descr sorts') dts;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   162
          val names = DatatypeProp.make_tnames Ts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   163
          val args = map Free (names ~~ Ts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   164
          val c = Const (cname, Ts ---> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   165
          fun perm_arg (dt, x) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   166
            let val T = type_of x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   167
            in if is_rec_type dt then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   168
                let val (Us, _) = strip_type T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   169
                in list_abs (map (pair "x") Us,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   170
                  Const (List.nth (perm_names_types, body_index dt)) $ pi $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   171
                    list_comb (x, map (fn (i, U) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   172
                      Const ("nominal.perm", permT --> U --> U) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   173
                        (Const ("List.rev", permT --> permT) $ pi) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   174
                        Bound i) ((length Us - 1 downto 0) ~~ Us)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   175
                end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   176
              else Const ("nominal.perm", permT --> T --> T) $ pi $ x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   177
            end;  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   178
        in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   179
          (("", HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   180
            (Const (List.nth (perm_names_types, i)) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   181
               Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   182
               list_comb (c, args),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   183
             list_comb (c, map perm_arg (dts ~~ args))))), [])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   184
        end) constrs
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   185
      end) descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   186
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   187
    val (thy2, perm_simps) = thy1 |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   188
      Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T, NoSyn))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   189
        (List.drop (perm_names_types, length new_type_names))) |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   190
      PrimrecPackage.add_primrec_i "" perm_eqs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   191
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   192
    (**** prove that permutation functions introduced by unfolding are ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   193
    (**** equivalent to already existing permutation functions         ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   194
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   195
    val _ = warning ("length descr: " ^ string_of_int (length descr));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   196
    val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   197
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   198
    val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   199
    val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   200
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   201
    val unfolded_perm_eq_thms =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   202
      if length descr = length new_type_names then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   203
      else map standard (List.drop (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   204
        (Goal.prove thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   205
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   206
            (map (fn (c as (s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   207
               let val [T1, T2] = binder_types T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   208
               in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   209
                 Const ("nominal.perm", T) $ pi $ Free (x, T2))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   210
               end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   211
             (perm_names_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   212
          (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   213
            ALLGOALS (asm_full_simp_tac
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   214
              (simpset_of thy2 addsimps [perm_fun_def]))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   215
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   216
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   217
    (**** prove [] \<bullet> t = t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   218
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   219
    val _ = warning "perm_empty_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   220
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   221
    val perm_empty_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   222
      let val permT = mk_permT (Type (a, []))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   223
      in map standard (List.take (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   224
        (Goal.prove thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   225
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   226
            (map (fn ((s, T), x) => HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   227
                (Const (s, permT --> T --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   228
                   Const ("List.list.Nil", permT) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   229
                 Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   230
             (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   231
              map body_type perm_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   232
          (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   233
            ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   234
        length new_type_names))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   235
      end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   236
      atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   237
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   238
    (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   239
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   240
    val _ = warning "perm_append_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   241
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   242
    (*FIXME: these should be looked up statically*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   243
    val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   244
    val pt2 = PureThy.get_thm thy2 (Name "pt2");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   245
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   246
    val perm_append_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   247
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   248
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   249
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   250
        val pi2 = Free ("pi2", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   251
        val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   252
        val pt2' = pt_inst RS pt2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   253
        val pt2_ax = PureThy.get_thm thy2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   254
          (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   255
      in List.take (map standard (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   256
        (Goal.prove thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   257
             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   258
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   259
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   260
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   261
                      (perm $ (Const ("List.op @", permT --> permT --> permT) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   262
                         pi1 $ pi2) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   263
                       perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   264
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   265
                  (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   266
                   map body_type perm_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   267
           (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   268
              ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   269
         length new_type_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   270
      end) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   271
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   272
    (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   273
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   274
    val _ = warning "perm_eq_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   275
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   276
    val pt3 = PureThy.get_thm thy2 (Name "pt3");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   277
    val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   278
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   279
    val perm_eq_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   280
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   281
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   282
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   283
        val pi2 = Free ("pi2", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   284
        (*FIXME: not robust - better access these theorems using NominalData?*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   285
        val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   286
        val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   287
        val pt3' = pt_inst RS pt3;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   288
        val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   289
        val pt3_ax = PureThy.get_thm thy2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   290
          (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   291
      in List.take (map standard (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   292
        (Goal.prove thy2 [] [] (Logic.mk_implies
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   293
             (HOLogic.mk_Trueprop (Const ("nominal.prm_eq",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   294
                permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   295
              HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   296
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   297
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   298
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   299
                      (perm $ pi1 $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   300
                       perm $ pi2 $ Free (x, T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   301
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   302
                  (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   303
                   map body_type perm_types ~~ perm_indnames)))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   304
           (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   305
              ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   306
         length new_type_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   307
      end) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   308
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   309
    (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   310
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   311
    val cp1 = PureThy.get_thm thy2 (Name "cp1");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   312
    val dj_cp = PureThy.get_thm thy2 (Name "dj_cp");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   313
    val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   314
    val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   315
    val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   316
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   317
    fun composition_instance name1 name2 thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   318
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   319
        val name1' = Sign.base_name name1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   320
        val name2' = Sign.base_name name2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   321
        val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   322
        val permT1 = mk_permT (Type (name1, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   323
        val permT2 = mk_permT (Type (name2, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   324
        val augment = map_type_tfree
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   325
          (fn (x, S) => TFree (x, cp_class :: S));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   326
        val Ts = map (augment o body_type) perm_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   327
        val cp_inst = PureThy.get_thm thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   328
          (Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   329
        val simps = simpset_of thy addsimps (perm_fun_def ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   330
          (if name1 <> name2 then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   331
             let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1'))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   332
             in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   333
           else
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   334
             let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   335
               val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   336
               val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst"))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   337
             in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   338
               [cp_inst RS cp1 RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   339
                at_inst RS (pt_inst RS pt_perm_compose) RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   340
                at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   341
            end))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   342
        val thms = split_conj_thm (standard (Goal.prove thy [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   343
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   344
              (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   345
                  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   346
                    val pi1 = Free ("pi1", permT1);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   347
                    val pi2 = Free ("pi2", permT2);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   348
                    val perm1 = Const (s, permT1 --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   349
                    val perm2 = Const (s, permT2 --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   350
                    val perm3 = Const ("nominal.perm", permT1 --> permT2 --> permT2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   351
                  in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   352
                    (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   353
                     perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   354
                  end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   355
                (perm_names ~~ Ts ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   356
          (fn _ => EVERY [indtac induction perm_indnames 1,
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   357
             ALLGOALS (asm_full_simp_tac simps)])))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   358
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   359
        foldl (fn ((s, tvs), thy) => AxClass.add_inst_arity_i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   360
            (s, replicate (length tvs) (cp_class :: classes), [cp_class])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   361
            (AxClass.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   362
          thy (full_new_type_names' ~~ tyvars)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   363
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   364
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   365
    val (thy3, perm_thmss) = thy2 |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   366
      fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   367
      curry (Library.foldr (fn ((i, (tyname, args, _)), thy) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   368
        AxClass.add_inst_arity_i (tyname, replicate (length args) classes, classes)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   369
        (AxClass.intro_classes_tac [] THEN REPEAT (EVERY
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   370
           [resolve_tac perm_empty_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   371
            resolve_tac perm_append_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   372
            resolve_tac perm_eq_thms 1, assume_tac 1])) thy))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   373
        (List.take (descr, length new_type_names)) |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   374
      PureThy.add_thmss
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   375
        [((space_implode "_" new_type_names ^ "_unfolded_perm_eq",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   376
          unfolded_perm_eq_thms), [Simplifier.simp_add_global]),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   377
         ((space_implode "_" new_type_names ^ "_perm_empty",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   378
          perm_empty_thms), [Simplifier.simp_add_global]),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   379
         ((space_implode "_" new_type_names ^ "_perm_append",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   380
          perm_append_thms), [Simplifier.simp_add_global]),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   381
         ((space_implode "_" new_type_names ^ "_perm_eq",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   382
          perm_eq_thms), [Simplifier.simp_add_global])];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   383
  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   384
    (**** Define representing sets ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   385
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   386
    val _ = warning "representing sets";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   387
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   388
    val rep_set_names = map (Sign.full_name thy3) (DatatypeProp.indexify_names
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   389
      (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   390
    val big_rep_name =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   391
      space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   392
        (fn (i, ("nominal.nOption", _, _)) => NONE
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   393
          | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   394
    val _ = warning ("big_rep_name: " ^ big_rep_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   395
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   396
    fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   397
          (case AList.lookup op = descr i of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   398
             SOME ("nominal.nOption", _, [(_, [dt']), _]) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   399
               apfst (cons dt) (strip_option dt')
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   400
           | _ => ([], dtf))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   401
      | strip_option dt = ([], dt);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   402
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   403
    fun make_intr s T (cname, cargs) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   404
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   405
        fun mk_prem (dt, (j, j', prems, ts)) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   406
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   407
            val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   408
            val (dts', dt'') = strip_dtyp dt';
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   409
            val Ts = map (typ_of_dtyp descr sorts') dts;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   410
            val Us = map (typ_of_dtyp descr sorts') dts';
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   411
            val T = typ_of_dtyp descr sorts' dt'';
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   412
            val free = mk_Free "x" (Us ---> T) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   413
            val free' = app_bnds free (length Us);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   414
            fun mk_abs_fun (T, (i, t)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   415
              let val U = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   416
              in (i + 1, Const ("nominal.abs_fun", [T, U, T] --->
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   417
                Type ("nominal.nOption", [U])) $ mk_Free "y" T i $ t)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   418
              end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   419
          in (j + 1, j' + length Ts,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   420
            case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   421
                DtRec k => list_all (map (pair "x") Us,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   422
                  HOLogic.mk_Trueprop (HOLogic.mk_mem (free',
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   423
                    Const (List.nth (rep_set_names, k),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   424
                      HOLogic.mk_setT T)))) :: prems
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   425
              | _ => prems,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   426
            snd (foldr mk_abs_fun (j', free) Ts) :: ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   427
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   428
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   429
        val (_, _, prems, ts) = foldr mk_prem (1, 1, [], []) cargs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   430
        val concl = HOLogic.mk_Trueprop (HOLogic.mk_mem
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   431
          (list_comb (Const (cname, map fastype_of ts ---> T), ts),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   432
           Const (s, HOLogic.mk_setT T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   433
      in Logic.list_implies (prems, concl)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   434
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   435
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   436
    val (intr_ts, ind_consts) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   437
      apfst List.concat (ListPair.unzip (List.mapPartial
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   438
        (fn ((_, ("nominal.nOption", _, _)), _) => NONE
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   439
          | ((i, (_, _, constrs)), rep_set_name) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   440
              let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   441
              in SOME (map (make_intr rep_set_name T) constrs,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   442
                Const (rep_set_name, HOLogic.mk_setT T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   443
              end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   444
                (descr ~~ rep_set_names)));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   445
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   446
    val (thy4, {raw_induct = rep_induct, intrs = rep_intrs, ...}) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   447
      setmp InductivePackage.quiet_mode false
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   448
        (InductivePackage.add_inductive_i false true big_rep_name false true false
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   449
           ind_consts (map (fn x => (("", x), [])) intr_ts) []) thy3;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   450
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   451
    (**** Prove that representing set is closed under permutation ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   452
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   453
    val _ = warning "proving closure under permutation...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   454
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   455
    val perm_indnames' = List.mapPartial
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   456
      (fn (x, (_, ("nominal.nOption", _, _))) => NONE | (x, _) => SOME x)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   457
      (perm_indnames ~~ descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   458
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   459
    fun mk_perm_closed name = map (fn th => standard (th RS mp))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   460
      (List.take (split_conj_thm (Goal.prove thy4 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   461
        (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   462
           (fn (S, x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   463
              let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   464
                val S = map_term_types (map_type_tfree
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   465
                  (fn (s, cs) => TFree (s, cs union cp_classes))) S;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   466
                val T = HOLogic.dest_setT (fastype_of S);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   467
                val permT = mk_permT (Type (name, []))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   468
              in HOLogic.mk_imp (HOLogic.mk_mem (Free (x, T), S),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   469
                HOLogic.mk_mem (Const ("nominal.perm", permT --> T --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   470
                  Free ("pi", permT) $ Free (x, T), S))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   471
              end) (ind_consts ~~ perm_indnames'))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   472
        (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
   473
           [indtac rep_induct [] 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   474
            ALLGOALS (simp_tac (simpset_of thy4 addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   475
              (symmetric perm_fun_def :: PureThy.get_thms thy4 (Name ("abs_perm"))))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   476
            ALLGOALS (resolve_tac rep_intrs 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   477
               THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   478
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   479
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   480
    (* FIXME: theorems are stored in database for testing only *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   481
    val perm_closed_thmss = map mk_perm_closed atoms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   482
    val (thy5, _) = PureThy.add_thmss [(("perm_closed",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   483
      List.concat perm_closed_thmss), [])] thy4;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   484
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   485
    (**** typedef ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   486
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   487
    val _ = warning "defining type...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   488
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   489
    val (thy6, typedefs) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   490
      foldl_map (fn (thy, ((((name, mx), tvs), c), name')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   491
        setmp TypedefPackage.quiet_mode true
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   492
          (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) c NONE
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   493
            (rtac exI 1 THEN
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   494
              QUIET_BREADTH_FIRST (has_fewer_prems 1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   495
              (resolve_tac rep_intrs 1))) thy |> (fn (thy, r) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   496
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   497
          val permT = mk_permT (TFree (variant tvs "'a", HOLogic.typeS));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   498
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   499
          val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   500
          val T = Type (Sign.intern_type thy name, tvs');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   501
          val Const (_, Type (_, [U])) = c
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   502
        in apsnd (pair r o hd)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   503
          (PureThy.add_defs_i true [(("prm_" ^ name ^ "_def", Logic.mk_equals
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   504
            (Const ("nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   505
             Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   506
               (Const ("nominal.perm", permT --> U --> U) $ pi $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   507
                 (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   508
                   Free ("x", T))))), [])] thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   509
        end))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   510
          (thy5, types_syntax ~~ tyvars ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   511
            (List.take (ind_consts, length new_type_names)) ~~ new_type_names);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   512
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   513
    val perm_defs = map snd typedefs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   514
    val Abs_inverse_thms = map (#Abs_inverse o fst) typedefs;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   515
    val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   516
    val Rep_thms = map (#Rep o fst) typedefs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   517
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   518
    val big_name = space_implode "_" new_type_names;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   519
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   520
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   521
    (** prove that new types are in class pt_<name> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   522
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   523
    val _ = warning "prove that new types are in class pt_<name> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   524
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   525
    fun pt_instance ((class, atom), perm_closed_thms) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   526
      fold (fn (((({Abs_inverse, Rep_inverse, Rep, ...},
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   527
        perm_def), name), tvs), perm_closed) => fn thy =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   528
          AxClass.add_inst_arity_i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   529
            (Sign.intern_type thy name,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   530
              replicate (length tvs) (classes @ cp_classes), [class])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   531
            (EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   532
              rewrite_goals_tac [perm_def],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   533
              asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   534
              asm_full_simp_tac (simpset_of thy addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   535
                [Rep RS perm_closed RS Abs_inverse]) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   536
              asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   537
                (Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   538
        (typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   539
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   540
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   541
    (** prove that new types are in class cp_<name1>_<name2> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   542
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   543
    val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   544
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   545
    fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   546
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   547
        val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   548
        val class = Sign.intern_class thy name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   549
        val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   550
      in fold (fn ((((({Abs_inverse, Rep_inverse, Rep, ...},
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   551
        perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   552
          AxClass.add_inst_arity_i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   553
            (Sign.intern_type thy name,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   554
              replicate (length tvs) (classes @ cp_classes), [class])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   555
            (EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   556
              rewrite_goals_tac [perm_def],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   557
              asm_full_simp_tac (simpset_of thy addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   558
                ((Rep RS perm_closed1 RS Abs_inverse) ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   559
                 (if atom1 = atom2 then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   560
                  else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   561
              cong_tac 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   562
              rtac refl 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   563
              rtac cp1' 1]) thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   564
        (typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms1 ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   565
          perm_closed_thms2) thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   566
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   567
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   568
    val thy7 = fold (fn x => fn thy => thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   569
      pt_instance x |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   570
      fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   571
        (classes ~~ atoms ~~ perm_closed_thmss) thy6;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   572
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   573
    (**** constructors ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   574
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   575
    fun mk_abs_fun (x, t) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   576
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   577
        val T = fastype_of x;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   578
        val U = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   579
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   580
        Const ("nominal.abs_fun", T --> U --> T -->
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   581
          Type ("nominal.nOption", [U])) $ x $ t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   582
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   583
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   584
    val (ty_idxs, _) = foldl
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   585
      (fn ((i, ("nominal.nOption", _, _)), p) => p
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   586
        | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   587
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   588
    fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   589
      | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   590
      | reindex dt = dt;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   591
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   592
    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
   593
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   594
    (** strips the "_Rep" in type names *)
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   595
    fun strip_nth_name i s = 
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   596
      let val xs = NameSpace.unpack s; 
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   597
      in NameSpace.pack (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   598
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   599
    val (descr'', ndescr) = ListPair.unzip (List.mapPartial
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   600
      (fn (i, ("nominal.nOption", _, _)) => NONE
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   601
        | (i, (s, dts, constrs)) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   602
             let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   603
               val SOME index = AList.lookup op = ty_idxs i;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   604
               val (constrs1, constrs2) = ListPair.unzip
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   605
                 (map (fn (cname, cargs) => apfst (pair (strip_nth_name 2 cname))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   606
                   (foldl_map (fn (dts, dt) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   607
                     let val (dts', dt') = strip_option dt
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   608
                     in (dts @ dts' @ [reindex dt'], (length dts, length dts')) end)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   609
                       ([], cargs))) constrs)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   610
             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
   611
               (index, constrs2))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   612
             end) descr);
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   613
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   614
    val (descr1, descr2) = splitAt (length new_type_names, descr'');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   615
    val descr' = [descr1, descr2];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   616
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   617
    val typ_of_dtyp' = replace_types' o typ_of_dtyp descr sorts';
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   618
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   619
    val rep_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   620
      Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   621
    val abs_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   622
      Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   623
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   624
    val recTs' = List.mapPartial
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   625
      (fn ((_, ("nominal.nOption", _, _)), T) => NONE
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   626
        | (_, T) => SOME T) (descr ~~ get_rec_types descr sorts');
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   627
    val recTs = get_rec_types descr'' sorts';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   628
    val newTs' = Library.take (length new_type_names, recTs');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   629
    val newTs = Library.take (length new_type_names, recTs);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   630
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   631
    val full_new_type_names = map (Sign.full_name (sign_of thy)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   632
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   633
    fun make_constr_def tname T T' ((thy, defs, eqns), ((cname, cargs), (cname', mx))) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   634
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   635
        fun constr_arg (dt, (j, l_args, r_args)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   636
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   637
            val x' = mk_Free "x" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   638
            val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   639
            val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp' dt) i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   640
              (dts ~~ (j upto j + length dts - 1))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   641
            val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   642
            val (dts', dt'') = strip_dtyp dt'
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   643
          in case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   644
              DtRec k => if k < length new_type_names then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   645
                  (j + length dts + 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   646
                   xs @ x :: l_args,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   647
                   foldr mk_abs_fun
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   648
                     (list_abs (map (pair "z" o typ_of_dtyp') dts',
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   649
                       Const (List.nth (rep_names, k), typ_of_dtyp' dt'' -->
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   650
                         typ_of_dtyp descr sorts' dt'') $ app_bnds x (length dts')))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   651
                     xs :: r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   652
                else error "nested recursion not (yet) supported"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   653
            | _ => (j + 1, x' :: l_args, x' :: r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   654
          end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   655
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   656
        val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   657
        val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   658
        val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   659
        val constrT = map fastype_of l_args ---> T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   660
        val lhs = list_comb (Const (Sign.full_name thy (Sign.base_name cname),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   661
          constrT), l_args);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   662
        val rhs = list_comb (Const (cname, map fastype_of r_args ---> T'), r_args);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   663
        val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   664
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   665
          (Const (rep_name, T --> T') $ lhs, rhs));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   666
        val def_name = (Sign.base_name cname) ^ "_def";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   667
        val (thy', [def_thm]) = thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   668
          Theory.add_consts_i [(cname', constrT, mx)] |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   669
          (PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   670
      in (thy', defs @ [def_thm], eqns @ [eqn]) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   671
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   672
    fun dt_constr_defs ((thy, defs, eqns, dist_lemmas),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   673
        (((((_, (_, _, constrs)), tname), T), T'), constr_syntax)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   674
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   675
        val rep_const = cterm_of thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   676
          (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   677
        val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   678
        val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   679
          ((Theory.add_path tname thy, defs, []), constrs ~~ constr_syntax)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   680
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   681
        (parent_path flat_names thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   682
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   683
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   684
    val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   685
      ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   686
        new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   687
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   688
    val abs_inject_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   689
      PureThy.get_thm thy8 (Name ("Abs_" ^ tname ^ "_inject"))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   690
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   691
    val rep_inject_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   692
      PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inject"))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   693
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   694
    val rep_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   695
      PureThy.get_thm thy8 (Name ("Rep_" ^ tname))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   696
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   697
    val rep_inverse_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   698
      PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inverse"))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   699
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   700
    (* prove theorem  Rep_i (Constr_j ...) = Constr'_j ...  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   701
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   702
    fun prove_constr_rep_thm eqn =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   703
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   704
        val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   705
        val rewrites = constr_defs @ map mk_meta_eq rep_inverse_thms
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   706
      in standard (Goal.prove thy8 [] [] eqn (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   707
        [resolve_tac inj_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   708
         rewrite_goals_tac rewrites,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   709
         rtac refl 3,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   710
         resolve_tac rep_intrs 2,
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   711
         REPEAT (resolve_tac rep_thms 1)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   712
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   713
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   714
    val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   715
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   716
    (* prove theorem  pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   717
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   718
    fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   719
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   720
        val _ $ (_ $ (Rep $ x) $ _) = Logic.unvarify (prop_of th);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   721
        val Type ("fun", [T, U]) = fastype_of Rep;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   722
        val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   723
        val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   724
      in
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   725
        standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   726
            (Const ("nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   727
             Rep $ (Const ("nominal.perm", permT --> T --> T) $ pi $ x))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   728
          (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   729
            perm_closed_thms @ Rep_thms)) 1))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   730
      end) Rep_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   731
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   732
    val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   733
      (atoms ~~ perm_closed_thmss));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   734
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   735
    (* prove distinctness theorems *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   736
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   737
    val distinct_props = setmp DatatypeProp.dtK 1000
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   738
      (DatatypeProp.make_distincts new_type_names descr' sorts') thy8;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   739
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   740
    val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   741
      dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   742
        (constr_rep_thmss ~~ dist_lemmas);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   743
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   744
    fun prove_distinct_thms (_, []) = []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   745
      | prove_distinct_thms (p as (rep_thms, dist_lemma), t::ts) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   746
          let
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   747
            val dist_thm = standard (Goal.prove thy8 [] [] t (fn _ =>
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   748
              simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   749
          in dist_thm::(standard (dist_thm RS not_sym))::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   750
            (prove_distinct_thms (p, ts))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   751
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   752
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   753
    val distinct_thms = map prove_distinct_thms
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   754
      (constr_rep_thmss ~~ dist_lemmas ~~ distinct_props);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   755
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   756
    (** prove equations for permutation functions **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   757
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   758
    val abs_perm = PureThy.get_thms thy8 (Name "abs_perm"); (* FIXME *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   759
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   760
    val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   761
      let val T = replace_types' (nth_dtyp i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   762
      in List.concat (map (fn (atom, perm_closed_thms) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   763
          map (fn ((cname, dts), constr_rep_thm) => 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   764
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   765
          val cname = Sign.intern_const thy8
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   766
            (NameSpace.append tname (Sign.base_name cname));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   767
          val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   768
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   769
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   770
          fun perm t =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   771
            let val T = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   772
            in Const ("nominal.perm", permT --> T --> T) $ pi $ t end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   773
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   774
          fun constr_arg (dt, (j, l_args, r_args)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   775
            let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   776
              val x' = mk_Free "x" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   777
              val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   778
              val Ts = map typ_of_dtyp' dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   779
              val xs = map (fn (T, i) => mk_Free "x" T i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   780
                (Ts ~~ (j upto j + length dts - 1))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   781
              val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   782
              val (dts', dt'') = strip_dtyp dt';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   783
            in case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   784
                DtRec k => if k < length new_type_names then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   785
                    (j + length dts + 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   786
                     xs @ x :: l_args,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   787
                     map perm (xs @ [x]) @ r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   788
                  else error "nested recursion not (yet) supported"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   789
              | _ => (j + 1, x' :: l_args, perm x' :: r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   790
            end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   791
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   792
          val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   793
          val c = Const (cname, map fastype_of l_args ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   794
        in
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   795
          standard (Goal.prove thy8 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   796
            (HOLogic.mk_Trueprop (HOLogic.mk_eq
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   797
              (perm (list_comb (c, l_args)), list_comb (c, r_args))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   798
            (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   799
              [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   800
               simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   801
                 constr_defs @ perm_closed_thms)) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   802
               TRY (simp_tac (HOL_basic_ss addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   803
                 (symmetric perm_fun_def :: abs_perm)) 1),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   804
               TRY (simp_tac (HOL_basic_ss addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   805
                 (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   806
                    perm_closed_thms)) 1)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   807
        end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   808
      end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   809
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   810
    (** prove injectivity of constructors **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   811
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   812
    val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   813
    val alpha = PureThy.get_thms thy8 (Name "alpha");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   814
    val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   815
    val fresh_def = PureThy.get_thm thy8 (Name "fresh_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   816
    val supp_def = PureThy.get_thm thy8 (Name "supp_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   817
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   818
    val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   819
      let val T = replace_types' (nth_dtyp i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   820
      in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   821
        if null dts then NONE else SOME
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   822
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   823
          val cname = Sign.intern_const thy8
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   824
            (NameSpace.append tname (Sign.base_name cname));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   825
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   826
          fun make_inj (dt, (j, args1, args2, eqs)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   827
            let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   828
              val x' = mk_Free "x" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   829
              val y' = mk_Free "y" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   830
              val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   831
              val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   832
              val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   833
              val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   834
              val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   835
              val y = mk_Free "y" (typ_of_dtyp' dt') (j + length dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   836
              val (dts', dt'') = strip_dtyp dt';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   837
            in case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   838
                DtRec k => if k < length new_type_names then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   839
                    (j + length dts + 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   840
                     xs @ (x :: args1), ys @ (y :: args2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   841
                     HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   842
                       (foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   843
                  else error "nested recursion not (yet) supported"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   844
              | _ => (j + 1, x' :: args1, y' :: args2, HOLogic.mk_eq (x', y') :: eqs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   845
            end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   846
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   847
          val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   848
          val Ts = map fastype_of args1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   849
          val c = Const (cname, Ts ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   850
        in
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   851
          standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   852
              (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   853
               foldr1 HOLogic.mk_conj eqs)))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   854
            (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   855
               [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   856
                  rep_inject_thms')) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   857
                TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   858
                  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
   859
                  perm_rep_perm_thms)) 1),
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
   860
                TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def ::
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   861
                  expand_fun_eq :: rep_inject_thms @ perm_rep_perm_thms)) 1)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   862
        end) (constrs ~~ constr_rep_thms)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   863
      end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   864
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   865
    (** equations for support and freshness **)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   866
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   867
    val Un_assoc = PureThy.get_thm thy8 (Name "Un_assoc");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   868
    val de_Morgan_conj = PureThy.get_thm thy8 (Name "de_Morgan_conj");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   869
    val Collect_disj_eq = PureThy.get_thm thy8 (Name "Collect_disj_eq");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   870
    val finite_Un = PureThy.get_thm thy8 (Name "finite_Un");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   871
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   872
    val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   873
      (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   874
      let val T = replace_types' (nth_dtyp i)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   875
      in List.concat (map (fn (cname, dts) => map (fn atom =>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   876
        let
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   877
          val cname = Sign.intern_const thy8
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   878
            (NameSpace.append tname (Sign.base_name cname));
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   879
          val atomT = Type (atom, []);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   880
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   881
          fun process_constr (dt, (j, args1, args2)) =
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   882
            let
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   883
              val x' = mk_Free "x" (typ_of_dtyp' dt) j;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   884
              val (dts, dt') = strip_option dt;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   885
              val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   886
              val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   887
              val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   888
              val (dts', dt'') = strip_dtyp dt';
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   889
            in case dt'' of
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   890
                DtRec k => if k < length new_type_names then
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   891
                    (j + length dts + 1,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   892
                     xs @ (x :: args1), foldr mk_abs_fun x xs :: args2)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   893
                  else error "nested recursion not (yet) supported"
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   894
              | _ => (j + 1, x' :: args1, x' :: args2)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   895
            end;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   896
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   897
          val (_, args1, args2) = foldr process_constr (1, [], []) dts;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   898
          val Ts = map fastype_of args1;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   899
          val c = list_comb (Const (cname, Ts ---> T), args1);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   900
          fun supp t =
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   901
            Const ("nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   902
          fun fresh t =
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   903
            Const ("nominal.fresh", atomT --> fastype_of t --> HOLogic.boolT) $
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   904
              Free ("a", atomT) $ t;
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   905
          val supp_thm = standard (Goal.prove thy8 [] []
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   906
              (HOLogic.mk_Trueprop (HOLogic.mk_eq
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   907
                (supp c,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   908
                 if null dts then Const ("{}", HOLogic.mk_setT atomT)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   909
                 else foldr1 (HOLogic.mk_binop "op Un") (map supp args2))))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   910
            (fn _ =>
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   911
              simp_tac (HOL_basic_ss addsimps (supp_def ::
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   912
                 Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
17874
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
   913
                 symmetric empty_def :: Finites.emptyI :: simp_thms @
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   914
                 abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   915
        in
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   916
          (supp_thm,
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   917
           standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   918
              (fresh c,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   919
               if null dts then HOLogic.true_const
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   920
               else foldr1 HOLogic.mk_conj (map fresh args2))))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   921
             (fn _ =>
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   922
               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
   923
        end) atoms) constrs)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   924
      end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   925
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   926
    (**** weak induction theorem ****)
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   927
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   928
    val arities = get_arities descr'';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   929
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   930
    fun mk_funs_inv thm =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   931
      let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   932
        val {sign, prop, ...} = rep_thm thm;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   933
        val _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ S)) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   934
          (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   935
        val used = add_term_tfree_names (a, []);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   936
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   937
        fun mk_thm i =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   938
          let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   939
            val Ts = map (TFree o rpair HOLogic.typeS)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   940
              (variantlist (replicate i "'t", used));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   941
            val f = Free ("f", Ts ---> U)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   942
          in standard (Goal.prove sign [] [] (Logic.mk_implies
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   943
            (HOLogic.mk_Trueprop (HOLogic.list_all
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   944
               (map (pair "x") Ts, HOLogic.mk_mem (app_bnds f i, S))),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   945
             HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   946
               r $ (a $ app_bnds f i)), f))))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   947
            (fn _ => EVERY [REPEAT (rtac ext 1), REPEAT (etac allE 1), rtac thm 1, atac 1]))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   948
          end
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   949
      in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   950
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   951
    fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   952
      let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   953
        val Rep_t = Const (List.nth (rep_names, i), T --> U) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   954
          mk_Free "x" T i;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   955
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   956
        val Abs_t =  Const (List.nth (abs_names, i), U --> T)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   957
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   958
      in (prems @ [HOLogic.imp $ HOLogic.mk_mem (Rep_t,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   959
            Const (List.nth (rep_set_names, i), HOLogic.mk_setT U)) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   960
              (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   961
          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
   962
      end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   963
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   964
    val (indrule_lemma_prems, indrule_lemma_concls) =
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   965
      Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   966
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   967
    val indrule_lemma = standard (Goal.prove thy8 [] []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   968
      (Logic.mk_implies
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   969
        (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   970
         HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   971
           [REPEAT (etac conjE 1),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   972
            REPEAT (EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   973
              [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   974
               etac mp 1, resolve_tac Rep_thms 1])]));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   975
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   976
    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
   977
    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
   978
      map (Free o apfst fst o dest_Var) Ps;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   979
    val indrule_lemma' = cterm_instantiate
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   980
      (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   981
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   982
    val Abs_inverse_thms' = List.concat (map mk_funs_inv Abs_inverse_thms);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   983
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   984
    val dt_induct_prop = DatatypeProp.make_ind descr' sorts';
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   985
    val dt_induct = standard (Goal.prove thy8 []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   986
      (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
   987
      (fn prems => EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   988
        [rtac indrule_lemma' 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   989
         (DatatypeAux.indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   990
         EVERY (map (fn (prem, r) => (EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   991
           [REPEAT (eresolve_tac Abs_inverse_thms' 1),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   992
            simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   993
            DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   994
                (prems ~~ constr_defs))]));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   995
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   996
    val case_names_induct = mk_case_names_induct descr'';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   997
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
   998
    (**** prove that new datatypes have finite support ****)
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
   999
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1000
    val _ = warning "proving finite support for the new datatype";
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1001
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1002
    val indnames = DatatypeProp.make_tnames recTs;
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1003
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1004
    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
  1005
    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
  1006
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1007
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1008
(*
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1009
    val finite_supp_thms = map (fn atom =>
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1010
      let val atomT = Type (atom, [])
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1011
      in map standard (List.take
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1012
        (split_conj_thm (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1013
           (foldr1 HOLogic.mk_conj (map (fn (s, T) => HOLogic.mk_mem
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1014
             (Const ("nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1015
              Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT atomT))))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1016
               (indnames ~~ recTs))))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1017
           (fn _ => indtac dt_induct indnames 1 THEN
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1018
            ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
  1019
              (abs_supp @ supp_atm @
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1020
               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
  1021
               List.concat supp_thms))))),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1022
         length new_type_names))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1023
      end) atoms;
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1024
*)
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1025
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1026
    (**** strong induction theorem ****)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1027
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1028
    val pnames = if length descr'' = 1 then ["P"]
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1029
      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
  1030
    val ind_sort = if null dt_atomTs then HOLogic.typeS
65e60434b3c2 Fixed problem with strong induction theorem for datatypes containing
berghofe
parents: 18142
diff changeset
  1031
      else map (fn T => Sign.intern_class thy8 ("fs_" ^
65e60434b3c2 Fixed problem with strong induction theorem for datatypes containing
berghofe
parents: 18142
diff changeset
  1032
        Sign.base_name (fst (dest_Type T)))) dt_atomTs;
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1033
    val fsT = TFree ("'n", ind_sort);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1034
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1035
    fun make_pred i T =
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1036
      Free (List.nth (pnames, i), T --> fsT --> HOLogic.boolT);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1037
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1038
    fun make_ind_prem k T ((cname, cargs), idxs) =
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1039
      let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1040
        val recs = List.filter is_rec_type cargs;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1041
        val Ts = map (typ_of_dtyp descr'' sorts') cargs;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1042
        val recTs' = map (typ_of_dtyp descr'' sorts') recs;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1043
        val tnames = variantlist (DatatypeProp.make_tnames Ts, pnames);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1044
        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
  1045
        val frees = tnames ~~ Ts;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1046
        val z = (variant tnames "z", fsT);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1047
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1048
        fun mk_prem ((dt, s), T) =
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1049
          let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1050
            val (Us, U) = strip_type T;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1051
            val l = length Us
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1052
          in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1053
            (make_pred (body_index dt) U $ app_bnds (Free (s, T)) l $ Bound l))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1054
          end;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1055
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1056
        val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1057
        val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1058
            (Const ("nominal.fresh", T --> fsT --> HOLogic.boolT) $
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1059
              Free p $ Free z))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1060
          (map (curry List.nth frees) (List.concat (map (fn (m, n) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1061
             m upto m + n - 1) idxs)))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1062
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1063
      in list_all_free (z :: frees, Logic.list_implies (prems' @ prems,
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1064
        HOLogic.mk_Trueprop (make_pred k T $ 
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1065
          list_comb (Const (cname, Ts ---> T), map Free frees) $ Free z)))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1066
      end;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1067
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1068
    val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1069
      map (make_ind_prem i T) (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1070
    val tnames = DatatypeProp.make_tnames recTs;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1071
    val z = (variant tnames "z", fsT);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1072
    val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1073
      (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T) $ Free z)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1074
        (descr'' ~~ recTs ~~ tnames)));
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1075
    val induct = Logic.list_implies (ind_prems, ind_concl);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1076
18104
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1077
    val simp_atts = replicate (length new_type_names) [Simplifier.simp_add_global];
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1078
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1079
    val (thy9, _) = thy8 |>
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1080
      Theory.add_path big_name |>
18017
f6abeac6dcb5 fixed case names in the weak induction principle and
urbanc
parents: 18016
diff changeset
  1081
      PureThy.add_thms [(("induct_weak", dt_induct), [case_names_induct])] |>>
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1082
      Theory.parent_path |>>>
18104
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1083
      DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms |>>>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1084
      DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss |>>>
18104
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1085
      DatatypeAux.store_thmss_atts "perm" new_type_names simp_atts perm_simps' |>>>
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1086
      DatatypeAux.store_thmss "inject" new_type_names inject_thms |>>>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1087
      DatatypeAux.store_thmss "supp" new_type_names supp_thms |>>>
18104
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1088
      DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms |>>
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1089
  (*    fold (fn (atom, ths) => fn thy =>
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1090
        let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom)
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1091
        in fold (fn T => AxClass.add_inst_arity_i
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1092
            (fst (dest_Type T),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1093
              replicate (length sorts) [class], [class])
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1094
            (AxClass.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1095
        end) (atoms ~~ finite_supp_thms) |>> *)
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1096
      Theory.add_path big_name |>>>
18142
a51ab4152097 called the induction principle "unsafe" instead of "test".
urbanc
parents: 18107
diff changeset
  1097
      PureThy.add_axioms_i [(("induct_unsafe", induct), [case_names_induct])] |>>
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1098
      Theory.parent_path;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1099
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1100
  in
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
  1101
    thy9
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1102
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1103
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1104
val add_nominal_datatype = gen_add_nominal_datatype read_typ true;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1105
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1106
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1107
(* FIXME: The following stuff should be exported by DatatypePackage *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1108
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1109
local structure P = OuterParse and K = OuterKeyword in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1110
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1111
val datatype_decl =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1112
  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1113
    (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1114
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1115
fun mk_datatype args =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1116
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1117
    val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1118
    val specs = map (fn ((((_, vs), t), mx), cons) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1119
      (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
  1120
  in add_nominal_datatype false names specs end;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1121
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1122
val nominal_datatypeP =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1123
  OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1124
    (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1125
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1126
val _ = OuterSyntax.add_parsers [nominal_datatypeP];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1127
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1128
end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1129
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1130
end