src/HOL/Nominal/nominal_datatype.ML
author wenzelm
Sat, 25 Jul 2015 23:41:53 +0200
changeset 60781 2da59cdf531c
parent 60754 02924903a6fd
child 60787 12cd198f07af
permissions -rw-r--r--
updated to infer_instantiate; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31936
9466169dc8e0 nominal.ML is nominal_datatype.ML
haftmann
parents: 31784
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_datatype.ML
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     2
    Author:     Stefan Berghofer and Christian Urban, TU Muenchen
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     3
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     4
Nominal datatype package for Isabelle/HOL.
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19489
diff changeset
     5
*)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     6
31936
9466169dc8e0 nominal.ML is nominal_datatype.ML
haftmann
parents: 31784
diff changeset
     7
signature NOMINAL_DATATYPE =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     8
sig
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
     9
  val nominal_datatype : Old_Datatype_Aux.config -> Old_Datatype.spec list -> theory -> theory
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    10
  val nominal_datatype_cmd : Old_Datatype_Aux.config -> Old_Datatype.spec_cmd list -> theory ->
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    11
    theory
22433
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
    12
  type descr
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    13
  type nominal_datatype_info
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    14
  val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    15
  val get_nominal_datatype : theory -> string -> nominal_datatype_info option
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
    16
  val mk_perm: typ list -> term -> term -> term
22529
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
    17
  val perm_of_pair: term * term -> term
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
    18
  val mk_not_sym: thm list -> thm list
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
    19
  val perm_simproc: simproc
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
    20
  val fresh_const: typ -> typ -> term
28373
5e2c526cfaf0 Added fresh_star_const.
berghofe
parents: 28084
diff changeset
    21
  val fresh_star_const: typ -> typ -> term
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    22
end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    23
31936
9466169dc8e0 nominal.ML is nominal_datatype.ML
haftmann
parents: 31784
diff changeset
    24
structure NominalDatatype : NOMINAL_DATATYPE =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    25
struct
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    26
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    27
val finite_emptyI = @{thm finite.emptyI};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    28
val finite_Diff = @{thm finite_Diff};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    29
val finite_Un = @{thm finite_Un};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    30
val Un_iff = @{thm Un_iff};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    31
val In0_eq = @{thm In0_eq};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    32
val In1_eq = @{thm In1_eq};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    33
val In0_not_In1 = @{thm In0_not_In1};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    34
val In1_not_In0 = @{thm In1_not_In0};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    35
val Un_assoc = @{thm Un_assoc};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    36
val Collect_disj_eq = @{thm Collect_disj_eq};
32129
d2aea34845d4 explicit antiquotation
haftmann
parents: 32124
diff changeset
    37
val Collect_False_empty = @{thm empty_def [THEN sym, THEN eq_reflection]};
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    38
val empty_iff = @{thm empty_iff};
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21540
diff changeset
    39
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents: 18067
diff changeset
    40
open NominalAtoms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    41
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    42
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22812
diff changeset
    43
(* theory data *)
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    44
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
    45
type descr =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    46
  (int * (string * Old_Datatype_Aux.dtyp list *
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
    47
      (string * (Old_Datatype_Aux.dtyp list * Old_Datatype_Aux.dtyp) list) list)) list;
22433
400fa18e951f - Changed format of descriptor contained in nominal_datatype_info
berghofe
parents: 22311
diff changeset
    48
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    49
type nominal_datatype_info =
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    50
  {index : int,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    51
   descr : descr,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    52
   rec_names : string list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    53
   rec_rewrites : thm list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    54
   induction : thm,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    55
   distinct : thm list,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    56
   inject : thm list};
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    57
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33368
diff changeset
    58
structure NominalDatatypesData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22812
diff changeset
    59
(
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    60
  type T = nominal_datatype_info Symtab.table;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    61
  val empty = Symtab.empty;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    62
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33368
diff changeset
    63
  fun merge data = Symtab.merge (K true) data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22812
diff changeset
    64
);
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    65
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    66
val get_nominal_datatypes = NominalDatatypesData.get;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    67
val put_nominal_datatypes = NominalDatatypesData.put;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    68
val map_nominal_datatypes = NominalDatatypesData.map;
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    69
val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22812
diff changeset
    70
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    71
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    72
(**** make datatype info ****)
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    73
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45741
diff changeset
    74
fun make_dt_info descr induct reccomb_names rec_thms
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33040
diff changeset
    75
    (i, (((_, (tname, _, _)), distinct), inject)) =
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    76
  (tname,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    77
   {index = i,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    78
    descr = descr,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    79
    rec_names = reccomb_names,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    80
    rec_rewrites = rec_thms,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    81
    induction = induct,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    82
    distinct = distinct,
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    83
    inject = inject});
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
    84
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    85
(*******************************)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
    86
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
    87
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (Thm.prems_of Old_Datatype.distinct_lemma);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    88
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    89
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
    90
(** simplification procedure for sorting permutations **)
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
    91
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
    92
val dj_cp = @{thm dj_cp};
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
    93
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
    94
fun dest_permT (Type (@{type_name fun},
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
    95
    [Type (@{type_name list}, [Type (@{type_name Product_Type.prod}, [T, _])]),
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
    96
      Type (@{type_name fun}, [_, U])])) = (T, U);
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
    97
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
    98
fun permTs_of (Const (@{const_name Nominal.perm}, T) $ t $ u) = fst (dest_permT T) :: permTs_of u
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
    99
  | permTs_of _ = [];
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   100
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   101
fun perm_simproc' ctxt (Const (@{const_name Nominal.perm}, T) $ t $ (u as Const (@{const_name Nominal.perm}, U) $ r $ s)) =
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   102
      let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   103
        val thy = Proof_Context.theory_of ctxt;
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   104
        val (aT as Type (a, []), S) = dest_permT T;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   105
        val (bT as Type (b, []), _) = dest_permT U;
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   106
      in if member (op =) (permTs_of u) aT andalso aT <> bT then
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   107
          let
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   108
            val cp = cp_inst_of thy a b;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   109
            val dj = dj_thm_of thy b a;
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   110
            val dj_cp' = [cp, dj] MRS dj_cp;
59635
025f70f35daf clarified context;
wenzelm
parents: 59621
diff changeset
   111
            val cert = SOME o Thm.cterm_of ctxt
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   112
          in
59635
025f70f35daf clarified context;
wenzelm
parents: 59621
diff changeset
   113
            SOME (mk_meta_eq (Drule.instantiate' [SOME (Thm.ctyp_of ctxt S)]
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   114
              [cert t, cert r, cert s] dj_cp'))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   115
          end
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   116
        else NONE
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   117
      end
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   118
  | perm_simproc' _ _ = NONE;
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   119
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   120
val perm_simproc =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38558
diff changeset
   121
  Simplifier.simproc_global @{theory} "perm_simp" ["pi1 \<bullet> (pi2 \<bullet> x)"] perm_simproc';
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
   122
60359
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
   123
fun projections ctxt rule =
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
   124
  Project_Rule.projections ctxt rule
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
   125
  |> map (Drule.export_without_context #> Rule_Cases.save rule);
18582
4f4cc426b440 provide projections of induct_weak, induct_unsafe;
wenzelm
parents: 18579
diff changeset
   126
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   127
val supp_prod = @{thm supp_prod};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   128
val fresh_prod = @{thm fresh_prod};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   129
val supports_fresh = @{thm supports_fresh};
44685
f5bc7d9d0d74 assert Pure equations for theorem references; avoid dynamic reference to fact
haftmann
parents: 44241
diff changeset
   130
val supports_def = Simpdata.mk_eq @{thm Nominal.supports_def};
f5bc7d9d0d74 assert Pure equations for theorem references; avoid dynamic reference to fact
haftmann
parents: 44241
diff changeset
   131
val fresh_def = Simpdata.mk_eq @{thm fresh_def};
f5bc7d9d0d74 assert Pure equations for theorem references; avoid dynamic reference to fact
haftmann
parents: 44241
diff changeset
   132
val supp_def = Simpdata.mk_eq @{thm supp_def};
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   133
val rev_simps = @{thms rev.simps};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   134
val app_simps = @{thms append.simps};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   135
val at_fin_set_supp = @{thm at_fin_set_supp};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   136
val at_fin_set_fresh = @{thm at_fin_set_fresh};
0dec18004e75 more antiquotations;
wenzelm
parents: 38864
diff changeset
   137
val abs_fun_eq1 = @{thm abs_fun_eq1};
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
   138
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   139
fun collect_simp ctxt = rewrite_rule ctxt [mk_meta_eq mem_Collect_eq];
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   140
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   141
fun mk_perm Ts t u =
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   142
  let
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   143
    val T = fastype_of1 (Ts, t);
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   144
    val U = fastype_of1 (Ts, u)
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   145
  in Const (@{const_name Nominal.perm}, T --> U --> U) $ t $ u end;
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
   146
22529
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   147
fun perm_of_pair (x, y) =
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   148
  let
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   149
    val T = fastype_of x;
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   150
    val pT = mk_permT T
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   151
  in Const (@{const_name Cons}, HOLogic.mk_prodT (T, T) --> pT --> pT) $
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   152
    HOLogic.mk_prod (x, y) $ Const (@{const_name Nil}, pT)
22529
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   153
  end;
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   154
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
   155
fun mk_not_sym ths = maps (fn th =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
   156
  (case Thm.prop_of th of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
   157
    _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
   158
    [th, th RS not_sym]
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
   159
  | _ => [th])) ths;
22529
902ed60d53a7 Exported perm_of_pair, mk_not_sym, and perm_simproc.
berghofe
parents: 22433
diff changeset
   160
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   161
fun fresh_const T U = Const (@{const_name Nominal.fresh}, T --> U --> HOLogic.boolT);
28373
5e2c526cfaf0 Added fresh_star_const.
berghofe
parents: 28084
diff changeset
   162
fun fresh_star_const T U =
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   163
  Const (@{const_name Nominal.fresh_star}, HOLogic.mk_setT T --> U --> HOLogic.boolT);
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
   164
59992
d8db5172c23f make SML/NJ more happy;
wenzelm
parents: 59936
diff changeset
   165
fun gen_nominal_datatype prep_specs (config: Old_Datatype_Aux.config) dts thy =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   166
  let
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   167
    val new_type_names = map (fn ((tname, _, _), _) => Binding.name_of tname) dts;
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   168
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   169
    val (dts', _) = prep_specs dts thy;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   170
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   171
    val atoms = atoms_of thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   172
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   173
    val tyvars = map (fn ((_, tvs, _), _) => tvs) dts';
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   174
    val sorts = flat tyvars;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   175
42388
a44b0fdaa6c2 standardized aliases of operations on tsig;
wenzelm
parents: 42375
diff changeset
   176
    fun inter_sort thy S S' = Sign.inter_sort thy (S, S');
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   177
    fun augment_sort_typ thy S =
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
   178
      let val S = Sign.minimize_sort thy (Sign.certify_sort thy S)
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   179
      in map_type_tfree (fn (s, S') => TFree (s,
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   180
        if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   181
      end;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   182
    fun augment_sort thy S = map_types (augment_sort_typ thy S);
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   183
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   184
    val types_syntax = map (fn ((tname, tvs, mx), constrs) => (tname, mx)) dts';
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   185
    val constr_syntax = map (fn (_, constrs) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   186
      map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   187
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   188
    val ps = map (fn ((n, _, _), _) =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   189
      (Sign.full_name thy n, Sign.full_name thy (Binding.suffix_name "_Rep" n))) dts;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   190
    val rps = map Library.swap ps;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   191
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   192
    fun replace_types (Type (@{type_name ABS}, [T, U])) =
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   193
          Type (@{type_name fun}, [T, Type (@{type_name noption}, [replace_types U])])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   194
      | replace_types (Type (s, Ts)) =
33035
15eab423e573 standardized basic operations on type option;
wenzelm
parents: 32960
diff changeset
   195
          Type (the_default s (AList.lookup op = ps s), map replace_types Ts)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   196
      | replace_types T = T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   197
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   198
    val dts'' = map (fn ((tname, tvs, mx), constrs) =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45838
diff changeset
   199
      ((Binding.suffix_name "_Rep" tname, tvs, NoSyn),
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   200
        map (fn (cname, cargs, mx) => (Binding.suffix_name "_Rep" cname,
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   201
          map replace_types cargs, NoSyn)) constrs)) dts';
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   202
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   203
    val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   204
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58300
diff changeset
   205
    val (full_new_type_names',thy1) = BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Kill_Type_Args] dts'' thy;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   206
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58300
diff changeset
   207
    val {descr, induct, ...} = BNF_LFP_Compat.the_info thy1 [] (hd full_new_type_names');
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   208
    fun nth_dtyp i = Old_Datatype_Aux.typ_of_dtyp descr (Old_Datatype_Aux.DtRec i);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   209
26651
e630c789ef2b overloading perm: use big_name;
wenzelm
parents: 26646
diff changeset
   210
    val big_name = space_implode "_" new_type_names;
e630c789ef2b overloading perm: use big_name;
wenzelm
parents: 26646
diff changeset
   211
e630c789ef2b overloading perm: use big_name;
wenzelm
parents: 26646
diff changeset
   212
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   213
    (**** define permutation functions ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   214
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   215
    val permT = mk_permT (TFree ("'x", @{sort type}));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   216
    val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   217
    val perm_types = map (fn (i, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   218
      let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   219
      in permT --> T --> T end) descr;
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   220
    val perm_names' = Old_Datatype_Prop.indexify_names (map (fn (i, _) =>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   221
      "perm_" ^ Old_Datatype_Aux.name_of_typ (nth_dtyp i)) descr);
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   222
    val perm_names = replicate (length new_type_names) @{const_name Nominal.perm} @
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
   223
      map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   224
    val perm_names_types = perm_names ~~ perm_types;
26680
18ff77116937 Adapted to new primrec package.
berghofe
parents: 26663
diff changeset
   225
    val perm_names_types' = perm_names' ~~ perm_types;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   226
26680
18ff77116937 Adapted to new primrec package.
berghofe
parents: 26663
diff changeset
   227
    val perm_eqs = maps (fn (i, (_, _, constrs)) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   228
      let val T = nth_dtyp i
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   229
      in map (fn (cname, dts) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   230
        let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   231
          val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) dts;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   232
          val names = Name.variant_list ["pi"] (Old_Datatype_Prop.make_tnames Ts);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   233
          val args = map Free (names ~~ Ts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   234
          val c = Const (cname, Ts ---> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   235
          fun perm_arg (dt, x) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   236
            let val T = type_of x
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   237
            in if Old_Datatype_Aux.is_rec_type dt then
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40627
diff changeset
   238
                let val Us = binder_types T
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   239
                in
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   240
                  fold_rev (Term.abs o pair "x") Us
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   241
                    (Free (nth perm_names_types' (Old_Datatype_Aux.body_index dt)) $ pi $
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   242
                      list_comb (x, map (fn (i, U) =>
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   243
                        Const (@{const_name Nominal.perm}, permT --> U --> U) $
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   244
                          (Const (@{const_name rev}, permT --> permT) $ pi) $
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   245
                          Bound i) ((length Us - 1 downto 0) ~~ Us)))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   246
                end
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   247
              else Const (@{const_name Nominal.perm}, permT --> T --> T) $ pi $ x
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   248
            end;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   249
        in
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
   250
          (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
26680
18ff77116937 Adapted to new primrec package.
berghofe
parents: 26663
diff changeset
   251
            (Free (nth perm_names_types' i) $
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   252
               Free ("pi", mk_permT (TFree ("'x", @{sort type}))) $
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   253
               list_comb (c, args),
26680
18ff77116937 Adapted to new primrec package.
berghofe
parents: 26663
diff changeset
   254
             list_comb (c, map perm_arg (dts ~~ args)))))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   255
        end) constrs
26680
18ff77116937 Adapted to new primrec package.
berghofe
parents: 26663
diff changeset
   256
      end) descr;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   257
26689
105031879f4a Auxiliary permutation functions are no longer declared using add_consts_i,
berghofe
parents: 26680
diff changeset
   258
    val (perm_simps, thy2) =
60003
ba8fa0c38d66 renamed ML funs
blanchet
parents: 59992
diff changeset
   259
      BNF_LFP_Compat.primrec_overloaded
26689
105031879f4a Auxiliary permutation functions are no longer declared using add_consts_i,
berghofe
parents: 26680
diff changeset
   260
        (map (fn (s, sT) => (s, sT, false))
105031879f4a Auxiliary permutation functions are no longer declared using add_consts_i,
berghofe
parents: 26680
diff changeset
   261
           (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   262
        (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   263
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   264
    (**** prove that permutation functions introduced by unfolding are ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   265
    (**** equivalent to already existing permutation functions         ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   266
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   267
    val _ = warning ("length descr: " ^ string_of_int (length descr));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   268
    val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   269
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   270
    val perm_indnames = Old_Datatype_Prop.make_tnames (map body_type perm_types);
44685
f5bc7d9d0d74 assert Pure equations for theorem references; avoid dynamic reference to fact
haftmann
parents: 44241
diff changeset
   271
    val perm_fun_def = Simpdata.mk_eq @{thm perm_fun_def};
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   272
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   273
    val unfolded_perm_eq_thms =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   274
      if length descr = length new_type_names then []
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   275
      else map Drule.export_without_context (List.drop (Old_Datatype_Aux.split_conj_thm
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   276
        (Goal.prove_global_future thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   277
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   278
            (map (fn (c as (s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   279
               let val [T1, T2] = binder_types T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   280
               in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   281
                 Const (@{const_name Nominal.perm}, T) $ pi $ Free (x, T2))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   282
               end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   283
             (perm_names_types ~~ perm_indnames))))
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   284
          (fn {context = ctxt, ...} =>
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   285
            EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1,
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   286
              ALLGOALS (asm_full_simp_tac (ctxt addsimps [perm_fun_def]))])),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   287
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   288
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   289
    (**** prove [] \<bullet> t = t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   290
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   291
    val _ = warning "perm_empty_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   292
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   293
    val perm_empty_thms = maps (fn a =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   294
      let val permT = mk_permT (Type (a, []))
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   295
      in map Drule.export_without_context (List.take (Old_Datatype_Aux.split_conj_thm
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   296
        (Goal.prove_global_future thy2 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   297
          (augment_sort thy2 [pt_class_of thy2 a]
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   298
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   299
              (map (fn ((s, T), x) => HOLogic.mk_eq
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   300
                  (Const (s, permT --> T --> T) $
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   301
                     Const (@{const_name Nil}, permT) $ Free (x, T),
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   302
                   Free (x, T)))
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   303
               (perm_names ~~
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   304
                map body_type perm_types ~~ perm_indnames)))))
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   305
          (fn {context = ctxt, ...} =>
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   306
            EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1,
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   307
              ALLGOALS (asm_full_simp_tac ctxt)])),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   308
        length new_type_names))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   309
      end)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   310
      atoms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   311
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   312
    (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   313
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   314
    val _ = warning "perm_append_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   315
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   316
    (*FIXME: these should be looked up statically*)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   317
    val at_pt_inst = Global_Theory.get_thm thy2 "at_pt_inst";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   318
    val pt2 = Global_Theory.get_thm thy2 "pt2";
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   319
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   320
    val perm_append_thms = maps (fn a =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   321
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   322
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   323
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   324
        val pi2 = Free ("pi2", permT);
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   325
        val pt_inst = pt_inst_of thy2 a;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   326
        val pt2' = pt_inst RS pt2;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   327
        val pt2_ax = Global_Theory.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   328
      in List.take (map Drule.export_without_context (Old_Datatype_Aux.split_conj_thm
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   329
        (Goal.prove_global_future thy2 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   330
           (augment_sort thy2 [pt_class_of thy2 a]
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   331
             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   332
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   333
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   334
                    in HOLogic.mk_eq
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   335
                      (perm $ (Const (@{const_name append}, permT --> permT --> permT) $
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   336
                         pi1 $ pi2) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   337
                       perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   338
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   339
                  (perm_names ~~
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   340
                   map body_type perm_types ~~ perm_indnames)))))
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   341
           (fn {context = ctxt, ...} =>
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   342
             EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1,
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   343
               ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt2', pt2_ax]))]))),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   344
         length new_type_names)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   345
      end) atoms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   346
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   347
    (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   348
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   349
    val _ = warning "perm_eq_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   350
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   351
    val pt3 = Global_Theory.get_thm thy2 "pt3";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   352
    val pt3_rev = Global_Theory.get_thm thy2 "pt3_rev";
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   353
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   354
    val perm_eq_thms = maps (fn a =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   355
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   356
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   357
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   358
        val pi2 = Free ("pi2", permT);
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   359
        val at_inst = at_inst_of thy2 a;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   360
        val pt_inst = pt_inst_of thy2 a;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   361
        val pt3' = pt_inst RS pt3;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   362
        val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   363
        val pt3_ax = Global_Theory.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   364
      in List.take (map Drule.export_without_context (Old_Datatype_Aux.split_conj_thm
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   365
        (Goal.prove_global_future thy2 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   366
          (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   367
             (HOLogic.mk_Trueprop (Const (@{const_name Nominal.prm_eq},
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   368
                permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   369
              HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   370
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   371
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   372
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   373
                      (perm $ pi1 $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   374
                       perm $ pi2 $ Free (x, T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   375
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   376
                  (perm_names ~~
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   377
                   map body_type perm_types ~~ perm_indnames))))))
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   378
           (fn {context = ctxt, ...} =>
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   379
             EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1,
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   380
               ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt3', pt3_rev', pt3_ax]))]))),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   381
         length new_type_names)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   382
      end) atoms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   383
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   384
    (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   385
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   386
    val cp1 = Global_Theory.get_thm thy2 "cp1";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   387
    val dj_cp = Global_Theory.get_thm thy2 "dj_cp";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   388
    val pt_perm_compose = Global_Theory.get_thm thy2 "pt_perm_compose";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   389
    val pt_perm_compose_rev = Global_Theory.get_thm thy2 "pt_perm_compose_rev";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   390
    val dj_perm_perm_forget = Global_Theory.get_thm thy2 "dj_perm_perm_forget";
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   391
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   392
    fun composition_instance name1 name2 thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   393
      let
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   394
        val cp_class = cp_class_of thy name1 name2;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   395
        val pt_class =
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   396
          if name1 = name2 then [pt_class_of thy name1]
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   397
          else [];
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   398
        val permT1 = mk_permT (Type (name1, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   399
        val permT2 = mk_permT (Type (name2, []));
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   400
        val Ts = map body_type perm_types;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   401
        val cp_inst = cp_inst_of thy name1 name2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   402
        fun simps ctxt = ctxt addsimps (perm_fun_def ::
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   403
          (if name1 <> name2 then
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   404
             let val dj = dj_thm_of thy name2 name1
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   405
             in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   406
           else
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   407
             let
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   408
               val at_inst = at_inst_of thy name1;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   409
               val pt_inst = pt_inst_of thy name1;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   410
             in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   411
               [cp_inst RS cp1 RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   412
                at_inst RS (pt_inst RS pt_perm_compose) RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   413
                at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   414
            end))
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
   415
        val sort = Sign.minimize_sort thy (Sign.certify_sort thy (cp_class :: pt_class));
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   416
        val thms = Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy [] []
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   417
          (augment_sort thy sort
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   418
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   419
              (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   420
                  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   421
                    val pi1 = Free ("pi1", permT1);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   422
                    val pi2 = Free ("pi2", permT2);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   423
                    val perm1 = Const (s, permT1 --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   424
                    val perm2 = Const (s, permT2 --> T --> T);
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   425
                    val perm3 = Const (@{const_name Nominal.perm}, permT1 --> permT2 --> permT2)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   426
                  in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   427
                    (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   428
                     perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   429
                  end)
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   430
                (perm_names ~~ Ts ~~ perm_indnames)))))
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   431
          (fn {context = ctxt, ...} =>
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   432
            EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1,
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   433
              ALLGOALS (asm_full_simp_tac (simps ctxt))]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   434
      in
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
   435
        fold (fn (s, tvs) => fn thy => Axclass.prove_arity
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   436
            (s, map (inter_sort thy sort o snd) tvs, [cp_class])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   437
            (fn ctxt' =>
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   438
              Class.intro_classes_tac ctxt' [] THEN ALLGOALS (resolve_tac ctxt' thms)) thy)
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   439
          (full_new_type_names' ~~ tyvars) thy
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   440
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   441
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   442
    val (perm_thmss,thy3) = thy2 |>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   443
      fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   444
      fold (fn atom => fn thy =>
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   445
        let val pt_name = pt_class_of thy atom
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   446
        in
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
   447
          fold (fn (s, tvs) => fn thy => Axclass.prove_arity
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   448
              (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58959
diff changeset
   449
              (fn ctxt => EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   450
                [Class.intro_classes_tac ctxt [],
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   451
                 resolve_tac ctxt perm_empty_thms 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   452
                 resolve_tac ctxt perm_append_thms 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   453
                 resolve_tac ctxt perm_eq_thms 1, assume_tac ctxt 1]) thy)
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   454
            (full_new_type_names' ~~ tyvars) thy
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   455
        end) atoms |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   456
      Global_Theory.add_thmss
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29389
diff changeset
   457
        [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   458
          unfolded_perm_eq_thms), [Simplifier.simp_add]),
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29389
diff changeset
   459
         ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   460
          perm_empty_thms), [Simplifier.simp_add]),
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29389
diff changeset
   461
         ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
   462
          perm_append_thms), [Simplifier.simp_add]),
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29389
diff changeset
   463
         ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   464
          perm_eq_thms), [Simplifier.simp_add])];
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   465
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   466
    (**** Define representing sets ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   467
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   468
    val _ = warning "representing sets";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   469
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   470
    val rep_set_names =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   471
      Old_Datatype_Prop.indexify_names
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   472
        (map (fn (i, _) => Old_Datatype_Aux.name_of_typ (nth_dtyp i) ^ "_set") descr);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   473
    val big_rep_name =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   474
      space_implode "_" (Old_Datatype_Prop.indexify_names (map_filter
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   475
        (fn (i, (@{type_name noption}, _, _)) => NONE
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   476
          | (i, _) => SOME (Old_Datatype_Aux.name_of_typ (nth_dtyp i))) descr)) ^ "_set";
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   477
    val _ = warning ("big_rep_name: " ^ big_rep_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   478
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   479
    fun strip_option (dtf as Old_Datatype_Aux.DtType ("fun", [dt, Old_Datatype_Aux.DtRec i])) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   480
          (case AList.lookup op = descr i of
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   481
             SOME (@{type_name noption}, _, [(_, [dt']), _]) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   482
               apfst (cons dt) (strip_option dt')
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   483
           | _ => ([], dtf))
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   484
      | strip_option (Old_Datatype_Aux.DtType ("fun",
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   485
            [dt, Old_Datatype_Aux.DtType (@{type_name noption}, [dt'])])) =
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   486
          apfst (cons dt) (strip_option dt')
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   487
      | strip_option dt = ([], dt);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   488
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   489
    val dt_atomTs = distinct op = (map (Old_Datatype_Aux.typ_of_dtyp descr)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   490
      (maps (fn (_, (_, _, cs)) => maps (maps (fst o strip_option) o snd) cs) descr));
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   491
    val dt_atoms = map (fst o dest_Type) dt_atomTs;
18280
45e139675daf Corrected atom class constraints in strong induction rule.
berghofe
parents: 18261
diff changeset
   492
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   493
    fun make_intr s T (cname, cargs) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   494
      let
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   495
        fun mk_prem dt (j, j', prems, ts) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   496
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   497
            val (dts, dt') = strip_option dt;
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   498
            val (dts', dt'') = Old_Datatype_Aux.strip_dtyp dt';
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   499
            val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) dts;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   500
            val Us = map (Old_Datatype_Aux.typ_of_dtyp descr) dts';
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   501
            val T = Old_Datatype_Aux.typ_of_dtyp descr dt'';
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   502
            val free = Old_Datatype_Aux.mk_Free "x" (Us ---> T) j;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   503
            val free' = Old_Datatype_Aux.app_bnds free (length Us);
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   504
            fun mk_abs_fun T (i, t) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   505
              let val U = fastype_of t
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   506
              in (i + 1, Const (@{const_name Nominal.abs_fun}, [T, U, T] --->
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   507
                Type (@{type_name noption}, [U])) $ Old_Datatype_Aux.mk_Free "y" T i $ t)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   508
              end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   509
          in (j + 1, j' + length Ts,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   510
            case dt'' of
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   511
                Old_Datatype_Aux.DtRec k => Logic.list_all (map (pair "x") Us,
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   512
                  HOLogic.mk_Trueprop (Free (nth rep_set_names k,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   513
                    T --> HOLogic.boolT) $ free')) :: prems
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   514
              | _ => prems,
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   515
            snd (fold_rev mk_abs_fun Ts (j', free)) :: ts)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   516
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   517
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   518
        val (_, _, prems, ts) = fold_rev mk_prem cargs (1, 1, [], []);
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   519
        val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   520
          list_comb (Const (cname, map fastype_of ts ---> T), ts))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   521
      in Logic.list_implies (prems, concl)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   522
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   523
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   524
    val (intr_ts, (rep_set_names', recTs')) =
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   525
      apfst flat (apsnd ListPair.unzip (ListPair.unzip (map_filter
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   526
        (fn ((_, (@{type_name noption}, _, _)), _) => NONE
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   527
          | ((i, (_, _, constrs)), rep_set_name) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   528
              let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   529
              in SOME (map (make_intr rep_set_name T) constrs,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   530
                (rep_set_name, T))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   531
              end)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   532
                (descr ~~ rep_set_names))));
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
   533
    val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   534
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   535
    val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   536
      thy3
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   537
      |> Sign.concealed
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33669
diff changeset
   538
      |> Inductive.add_inductive_global
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
   539
          {quiet_mode = false, verbose = false, alt_name = Binding.name big_rep_name,
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 46961
diff changeset
   540
           coind = false, no_elim = true, no_ind = false, skip_mono = true}
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
   541
          (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
26128
fe2d24c26e0c inductive package: simplified group handling;
wenzelm
parents: 26067
diff changeset
   542
             (rep_set_names' ~~ recTs'))
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   543
          [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) []
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   544
      ||> Sign.restore_naming thy3;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   545
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   546
    (**** Prove that representing set is closed under permutation ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   547
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   548
    val _ = warning "proving closure under permutation...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   549
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   550
    val abs_perm = Global_Theory.get_thms thy4 "abs_perm";
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
   551
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   552
    val perm_indnames' = map_filter
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   553
      (fn (x, (_, (@{type_name noption}, _, _))) => NONE | (x, _) => SOME x)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   554
      (perm_indnames ~~ descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   555
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
   556
    fun mk_perm_closed name = map (fn th => Drule.export_without_context (th RS mp))
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   557
      (List.take (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy4 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   558
        (augment_sort thy4
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33035
diff changeset
   559
          (pt_class_of thy4 name :: map (cp_class_of thy4 name) (remove (op =) name dt_atoms))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   560
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   561
            (fn ((s, T), x) =>
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   562
               let
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   563
                 val S = Const (s, T --> HOLogic.boolT);
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   564
                 val permT = mk_permT (Type (name, []))
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   565
               in HOLogic.mk_imp (S $ Free (x, T),
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   566
                 S $ (Const (@{const_name Nominal.perm}, permT --> T --> T) $
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   567
                   Free ("pi", permT) $ Free (x, T)))
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   568
               end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
51671
0d142a78fb7c formal proof context for axclass proofs;
wenzelm
parents: 51122
diff changeset
   569
        (fn {context = ctxt, ...} => EVERY
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
   570
           [Old_Datatype_Aux.ind_tac ctxt rep_induct [] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   571
            ALLGOALS (simp_tac (ctxt addsimps
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   572
              (Thm.symmetric perm_fun_def :: abs_perm))),
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   573
            ALLGOALS (resolve_tac ctxt rep_intrs THEN_ALL_NEW assume_tac ctxt)])),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   574
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   575
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   576
    val perm_closed_thmss = map mk_perm_closed atoms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   577
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   578
    (**** typedef ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   579
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   580
    val _ = warning "defining type...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   581
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   582
    val (typedefs, thy6) =
26651
e630c789ef2b overloading perm: use big_name;
wenzelm
parents: 26646
diff changeset
   583
      thy4
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   584
      |> fold_map (fn (((name, mx), tvs), (cname, U)) => fn thy =>
58239
1c5bc387bd4c added flag to 'typedef' to allow concealed definitions
blanchet
parents: 58238
diff changeset
   585
          Typedef.add_typedef_global false
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   586
            (name, map (fn (v, _) => (v, dummyS)) tvs, mx)  (* FIXME keep constraints!? *)
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37391
diff changeset
   587
            (Const (@{const_name Collect}, (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   588
               Const (cname, U --> HOLogic.boolT)) NONE
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   589
            (fn ctxt =>
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   590
              resolve_tac ctxt [exI] 1 THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   591
              resolve_tac ctxt [CollectI] 1 THEN
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   592
              QUIET_BREADTH_FIRST (has_fewer_prems 1)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   593
              (resolve_tac ctxt rep_intrs 1)) thy |> (fn ((_, r), thy) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   594
        let
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   595
          val permT = mk_permT
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   596
            (TFree (singleton (Name.variant_list (map fst tvs)) "'a", @{sort type}));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   597
          val pi = Free ("pi", permT);
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   598
          val T = Type (Sign.full_name thy name, map TFree tvs);
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   599
        in apfst (pair r o hd)
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   600
          (Global_Theory.add_defs_unchecked true
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   601
            [((Binding.map_name (fn n => "prm_" ^ n ^ "_def") name, Logic.mk_equals
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   602
              (Const (@{const_name Nominal.perm}, permT --> T --> T) $ pi $ Free ("x", T),
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   603
               Const (Sign.intern_const thy ("Abs_" ^ Binding.name_of name), U --> T) $
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   604
                 (Const (@{const_name Nominal.perm}, permT --> U --> U) $ pi $
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   605
                   (Const (Sign.intern_const thy ("Rep_" ^ Binding.name_of name), T --> U) $
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
   606
                     Free ("x", T))))), [])] thy)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   607
        end))
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   608
        (types_syntax ~~ tyvars ~~ List.take (rep_set_names'' ~~ recTs', length new_type_names));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   609
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   610
    val ctxt6 = Proof_Context.init_global thy6;
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   611
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   612
    val perm_defs = map snd typedefs;
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   613
    val Abs_inverse_thms = map (collect_simp ctxt6 o #Abs_inverse o snd o fst) typedefs;
35994
9cc3df9a606e Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents: 35845
diff changeset
   614
    val Rep_inverse_thms = map (#Rep_inverse o snd o fst) typedefs;
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   615
    val Rep_thms = map (collect_simp ctxt6 o #Rep o snd o fst) typedefs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   616
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   617
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   618
    (** prove that new types are in class pt_<name> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   619
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   620
    val _ = warning "prove that new types are in class pt_<name> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   621
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   622
    fun pt_instance (atom, perm_closed_thms) =
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   623
      fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   624
        perm_def), name), tvs), perm_closed) => fn thy =>
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   625
          let
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   626
            val pt_class = pt_class_of thy atom;
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
   627
            val sort = Sign.minimize_sort thy (Sign.certify_sort thy
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
   628
              (pt_class :: map (cp_class_of thy atom) (remove (op =) atom dt_atoms)))
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
   629
          in Axclass.prove_arity
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   630
            (Sign.intern_type thy name,
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   631
              map (inter_sort thy sort o snd) tvs, [pt_class])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   632
            (fn ctxt => EVERY [Class.intro_classes_tac ctxt [],
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   633
              rewrite_goals_tac ctxt [perm_def],
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   634
              asm_full_simp_tac (ctxt addsimps [Rep_inverse]) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   635
              asm_full_simp_tac (ctxt addsimps
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   636
                [Rep RS perm_closed RS Abs_inverse]) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   637
              asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Global_Theory.get_thm thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   638
                ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   639
          end)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   640
        (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   641
           new_type_names ~~ tyvars ~~ perm_closed_thms);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   642
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   643
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   644
    (** prove that new types are in class cp_<name1>_<name2> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   645
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   646
    val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   647
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   648
    fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   649
      let
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   650
        val cp_class = cp_class_of thy atom1 atom2;
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
   651
        val sort = Sign.minimize_sort thy (Sign.certify_sort thy
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33035
diff changeset
   652
          (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (remove (op =) atom1 dt_atoms) @
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   653
           (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
   654
            pt_class_of thy atom2 :: map (cp_class_of thy atom2) (remove (op =) atom2 dt_atoms))));
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   655
        val cp1' = cp_inst_of thy atom1 atom2 RS cp1
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   656
      in fold (fn ((((((Abs_inverse, Rep),
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   657
        perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
   658
          Axclass.prove_arity
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   659
            (Sign.intern_type thy name,
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
   660
              map (inter_sort thy sort o snd) tvs, [cp_class])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   661
            (fn ctxt => EVERY [Class.intro_classes_tac ctxt [],
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   662
              rewrite_goals_tac ctxt [perm_def],
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   663
              asm_full_simp_tac (ctxt addsimps
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   664
                ((Rep RS perm_closed1 RS Abs_inverse) ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   665
                 (if atom1 = atom2 then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   666
                  else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 58372
diff changeset
   667
              cong_tac ctxt 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   668
              resolve_tac ctxt [refl] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   669
              resolve_tac ctxt [cp1'] 1]) thy)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   670
        (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   671
           tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   672
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   673
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   674
    val thy7 = fold (fn x => fn thy => thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   675
      pt_instance x |>
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   676
      fold (cp_instance x) (atoms ~~ perm_closed_thmss))
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   677
        (atoms ~~ perm_closed_thmss) thy6;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   678
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   679
    (**** constructors ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   680
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   681
    fun mk_abs_fun x t =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   682
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   683
        val T = fastype_of x;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   684
        val U = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   685
      in
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   686
        Const (@{const_name Nominal.abs_fun}, T --> U --> T -->
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   687
          Type (@{type_name noption}, [U])) $ x $ t
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   688
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   689
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29585
diff changeset
   690
    val (ty_idxs, _) = List.foldl
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   691
      (fn ((i, (@{type_name noption}, _, _)), p) => p
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   692
        | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   693
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   694
    fun reindex (Old_Datatype_Aux.DtType (s, dts)) = Old_Datatype_Aux.DtType (s, map reindex dts)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   695
      | reindex (Old_Datatype_Aux.DtRec i) =
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   696
        Old_Datatype_Aux.DtRec (the (AList.lookup op = ty_idxs i))
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   697
      | reindex dt = dt;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   698
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 39557
diff changeset
   699
    fun strip_suffix i s = implode (List.take (raw_explode s, size s - i));  (* FIXME Symbol.explode (?) *)
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   700
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   701
    (** strips the "_Rep" in type names *)
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   702
    fun strip_nth_name i s =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   703
      let val xs = Long_Name.explode s;
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   704
      in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   705
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   706
    val (descr'', ndescr) = ListPair.unzip (map_filter
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   707
      (fn (i, (@{type_name noption}, _, _)) => NONE
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   708
        | (i, (s, dts, constrs)) =>
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   709
             let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   710
               val SOME index = AList.lookup op = ty_idxs i;
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   711
               val (constrs2, constrs1) =
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   712
                 map_split (fn (cname, cargs) =>
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   713
                   apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   714
                   (fold_map (fn dt => fn dts =>
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   715
                     let val (dts', dt') = strip_option dt
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   716
                     in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end)
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 30450
diff changeset
   717
                       cargs [])) constrs
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   718
             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
   719
               (index, constrs2))
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
   720
             end) descr);
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
   721
19489
4441b637871b SplitAt -> chop
berghofe
parents: 19403
diff changeset
   722
    val (descr1, descr2) = chop (length new_type_names) descr'';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   723
    val descr' = [descr1, descr2];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   724
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
   725
    fun partition_cargs idxs xs = map (fn (i, j) =>
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   726
      (List.take (List.drop (xs, i), j), nth xs (i + j))) idxs;
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
   727
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   728
    val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   729
      map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   730
        (constrs ~~ idxss)))) (descr'' ~~ ndescr);
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   731
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   732
    fun nth_dtyp' i = Old_Datatype_Aux.typ_of_dtyp descr'' (Old_Datatype_Aux.DtRec i);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   733
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   734
    val rep_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   735
      Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   736
    val abs_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   737
      Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   738
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   739
    val recTs = Old_Datatype_Aux.get_rec_types descr'';
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   740
    val newTs' = take (length new_type_names) recTs';
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   741
    val newTs = take (length new_type_names) recTs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   742
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
   743
    val full_new_type_names = map (Sign.full_bname thy) new_type_names;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   744
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   745
    fun make_constr_def tname T T' (((cname_rep, _), (cname, cargs)), (cname', mx))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   746
        (thy, defs, eqns) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   747
      let
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   748
        fun constr_arg (dts, dt) (j, l_args, r_args) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   749
          let
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   750
            val xs =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   751
              map (fn (dt, i) =>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   752
                  Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) i)
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   753
                (dts ~~ (j upto j + length dts - 1))
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   754
            val x =
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   755
              Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   756
                (j + length dts)
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   757
          in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   758
            (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   759
             xs @ x :: l_args,
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   760
             fold_rev mk_abs_fun xs
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   761
               (case dt of
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   762
                  Old_Datatype_Aux.DtRec k => if k < length new_type_names then
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   763
                      Const (nth rep_names k, Old_Datatype_Aux.typ_of_dtyp descr'' dt -->
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   764
                        Old_Datatype_Aux.typ_of_dtyp descr dt) $ x
58372
bfd497f2f4c2 moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents: 58354
diff changeset
   765
                    else error "nested recursion not supported"
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   766
                | _ => x) :: r_args)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   767
          end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   768
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   769
        val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []);
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22543
diff changeset
   770
        val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22543
diff changeset
   771
        val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   772
        val constrT = map fastype_of l_args ---> T;
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   773
        val lhs = list_comb (Const (cname, constrT), l_args);
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   774
        val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   775
        val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   776
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   777
          (Const (rep_name, T --> T') $ lhs, rhs));
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   778
        val def_name = (Long_Name.base_name cname) ^ "_def";
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18350
diff changeset
   779
        val ([def_thm], thy') = thy |>
56239
17df7145a871 tuned signature;
wenzelm
parents: 55990
diff changeset
   780
          Sign.add_consts [(cname', constrT, mx)] |>
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   781
          (Global_Theory.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   782
      in (thy', defs @ [def_thm], eqns @ [eqn]) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   783
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   784
    fun dt_constr_defs ((((((_, (_, _, constrs)),
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   785
        (_, (_, _, constrs'))), tname), T), T'), constr_syntax) (thy, defs, eqns, dist_lemmas) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   786
      let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   787
        val rep_const = Thm.global_cterm_of thy
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   788
          (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   789
        val dist =
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
   790
          Drule.export_without_context
60781
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   791
            (infer_instantiate (Proof_Context.init_global thy)
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
   792
              [(#1 (dest_Var distinct_f), rep_const)] Old_Datatype.distinct_lemma);
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   793
        val (thy', defs', eqns') = fold (make_constr_def tname T T')
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   794
          (constrs ~~ constrs' ~~ constr_syntax) (Sign.add_path tname thy, defs, [])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   795
      in
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   796
        (Sign.parent_path thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   797
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   798
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   799
    val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = fold dt_constr_defs
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   800
      (List.take (descr, length new_type_names) ~~
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   801
        List.take (pdescr, length new_type_names) ~~
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   802
        new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax)
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   803
      (thy7, [], [], []);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   804
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   805
    val abs_inject_thms = map (collect_simp ctxt6 o #Abs_inject o snd o fst) typedefs
35994
9cc3df9a606e Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents: 35845
diff changeset
   806
    val rep_inject_thms = map (#Rep_inject o snd o fst) typedefs
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   807
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   808
    (* prove theorem  Rep_i (Constr_j ...) = Constr'_j ...  *)
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   809
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   810
    fun prove_constr_rep_thm eqn =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   811
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   812
        val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
   813
        val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   814
      in Goal.prove_global_future thy8 [] [] eqn (fn {context = ctxt, ...} => EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   815
        [resolve_tac ctxt inj_thms 1,
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   816
         rewrite_goals_tac ctxt rewrites,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   817
         resolve_tac ctxt [refl] 3,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   818
         resolve_tac ctxt rep_intrs 2,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
   819
         REPEAT (resolve_tac ctxt Rep_thms 1)])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   820
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   821
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   822
    val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   823
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   824
    (* prove theorem  pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   825
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   826
    fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   827
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
   828
        val _ $ (_ $ (Rep $ x)) = Logic.unvarify_global (Thm.prop_of th);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   829
        val Type ("fun", [T, U]) = fastype_of Rep;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   830
        val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   831
        val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   832
      in
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   833
        Goal.prove_global_future thy8 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   834
          (augment_sort thy8
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33035
diff changeset
   835
            (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (remove (op =) atom dt_atoms))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   836
            (HOLogic.mk_Trueprop (HOLogic.mk_eq
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   837
              (Const (@{const_name Nominal.perm}, permT --> U --> U) $ pi $ (Rep $ x),
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   838
               Rep $ (Const (@{const_name Nominal.perm}, permT --> T --> T) $ pi $ x)))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   839
          (fn {context = ctxt, ...} =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   840
            simp_tac (put_simpset HOL_basic_ss ctxt addsimps (perm_defs @ Abs_inverse_thms @
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   841
            perm_closed_thms @ Rep_thms)) 1)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   842
      end) Rep_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   843
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   844
    val perm_rep_perm_thms = maps prove_perm_rep_perm (atoms ~~ perm_closed_thmss);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   845
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   846
    (* prove distinctness theorems *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   847
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   848
    val distinct_props = Old_Datatype_Prop.make_distincts descr';
27300
4cb3101d2bf7 DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
haftmann
parents: 27275
diff changeset
   849
    val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
4cb3101d2bf7 DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
haftmann
parents: 27275
diff changeset
   850
      dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
4cb3101d2bf7 DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
haftmann
parents: 27275
diff changeset
   851
        constr_rep_thmss dist_lemmas;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   852
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   853
    fun prove_distinct_thms _ [] = []
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   854
      | prove_distinct_thms (p as (rep_thms, dist_lemma)) (t :: ts) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   855
          let
51671
0d142a78fb7c formal proof context for axclass proofs;
wenzelm
parents: 51122
diff changeset
   856
            val dist_thm = Goal.prove_global_future thy8 [] [] t (fn {context = ctxt, ...} =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   857
              simp_tac (ctxt addsimps (dist_lemma :: rep_thms)) 1)
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
   858
          in
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
   859
            dist_thm :: Drule.export_without_context (dist_thm RS not_sym) ::
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   860
              prove_distinct_thms p ts
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   861
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   862
26969
cf3f998d0631 moved case distinction over number of constructors for distinctness rules from DatatypeProp to DatatypeRepProofs
haftmann
parents: 26966
diff changeset
   863
    val distinct_thms = map2 prove_distinct_thms
cf3f998d0631 moved case distinction over number of constructors for distinctness rules from DatatypeProp to DatatypeRepProofs
haftmann
parents: 26966
diff changeset
   864
      (constr_rep_thmss ~~ dist_lemmas) distinct_props;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   865
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   866
    (** prove equations for permutation functions **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   867
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   868
    val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   869
      let val T = nth_dtyp' i
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   870
      in maps (fn (atom, perm_closed_thms) =>
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   871
          map (fn ((cname, dts), constr_rep_thm) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   872
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   873
          val cname = Sign.intern_const thy8
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   874
            (Long_Name.append tname (Long_Name.base_name cname));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   875
          val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   876
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   877
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   878
          fun perm t =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   879
            let val T = fastype_of t
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
   880
            in Const (@{const_name Nominal.perm}, permT --> T --> T) $ pi $ t end;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   881
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   882
          fun constr_arg (dts, dt) (j, l_args, r_args) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   883
            let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   884
              val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') dts;
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   885
              val xs =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   886
                map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i)
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   887
                  (Ts ~~ (j upto j + length dts - 1));
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   888
              val x =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   889
                Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   890
                  (j + length dts);
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   891
            in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   892
              (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   893
               xs @ x :: l_args,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   894
               map perm (xs @ [x]) @ r_args)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   895
            end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   896
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   897
          val (_, l_args, r_args) = fold_rev constr_arg dts (1, [], []);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   898
          val c = Const (cname, map fastype_of l_args ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   899
        in
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   900
          Goal.prove_global_future thy8 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   901
            (augment_sort thy8
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33035
diff changeset
   902
              (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (remove (op =) atom dt_atoms))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   903
              (HOLogic.mk_Trueprop (HOLogic.mk_eq
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   904
                (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
51671
0d142a78fb7c formal proof context for axclass proofs;
wenzelm
parents: 51122
diff changeset
   905
            (fn {context = ctxt, ...} => EVERY
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   906
              [simp_tac (ctxt addsimps (constr_rep_thm :: perm_defs)) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   907
               simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Rep_thms @ Abs_inverse_thms @
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   908
                 constr_defs @ perm_closed_thms)) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   909
               TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   910
                 (Thm.symmetric perm_fun_def :: abs_perm)) 1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   911
               TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   912
                 (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
   913
                    perm_closed_thms)) 1)])
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   914
        end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss)
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   915
      end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   916
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   917
    (** prove injectivity of constructors **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   918
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   919
    val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   920
    val alpha = Global_Theory.get_thms thy8 "alpha";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   921
    val abs_fresh = Global_Theory.get_thms thy8 "abs_fresh";
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   922
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   923
    val pt_cp_sort =
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   924
      map (pt_class_of thy8) dt_atoms @
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33035
diff changeset
   925
      maps (fn s => map (cp_class_of thy8 s) (remove (op =) s dt_atoms)) dt_atoms;
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   926
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   927
    val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   928
      let val T = nth_dtyp' i
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   929
      in map_filter (fn ((cname, dts), constr_rep_thm) =>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   930
        if null dts then NONE else SOME
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   931
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   932
          val cname = Sign.intern_const thy8
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   933
            (Long_Name.append tname (Long_Name.base_name cname));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   934
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   935
          fun make_inj (dts, dt) (j, args1, args2, eqs) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   936
            let
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   937
              val Ts_idx =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   938
                map (Old_Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1);
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   939
              val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) Ts_idx;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   940
              val ys = map (fn (T, i) => Old_Datatype_Aux.mk_Free "y" T i) Ts_idx;
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   941
              val x =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   942
                Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   943
                  (j + length dts);
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   944
              val y =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   945
                Old_Datatype_Aux.mk_Free "y" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   946
                  (j + length dts);
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   947
            in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   948
              (j + length dts + 1,
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   949
               xs @ (x :: args1), ys @ (y :: args2),
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   950
               HOLogic.mk_eq
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   951
                 (fold_rev mk_abs_fun xs x, fold_rev mk_abs_fun ys y) :: eqs)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   952
            end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   953
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   954
          val (_, args1, args2, eqs) = fold_rev make_inj dts (1, [], [], []);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   955
          val Ts = map fastype_of args1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   956
          val c = Const (cname, Ts ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   957
        in
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
   958
          Goal.prove_global_future thy8 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   959
            (augment_sort thy8 pt_cp_sort
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   960
              (HOLogic.mk_Trueprop (HOLogic.mk_eq
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   961
                (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   962
                 foldr1 HOLogic.mk_conj eqs))))
51671
0d142a78fb7c formal proof context for axclass proofs;
wenzelm
parents: 51122
diff changeset
   963
            (fn {context = ctxt, ...} => EVERY
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   964
               [asm_full_simp_tac (ctxt addsimps (constr_rep_thm ::
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   965
                  rep_inject_thms')) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   966
                TRY (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   967
                  addsimps (fresh_def :: supp_def ::
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   968
                  alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
   969
                  perm_rep_perm_thms)) 1)])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   970
        end) (constrs ~~ constr_rep_thms)
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   971
      end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   972
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   973
    (** equations for support and freshness **)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   974
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   975
    val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   976
      (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
   977
      let val T = nth_dtyp' i
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
   978
      in maps (fn (cname, dts) => map (fn atom =>
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   979
        let
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   980
          val cname = Sign.intern_const thy8
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   981
            (Long_Name.append tname (Long_Name.base_name cname));
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   982
          val atomT = Type (atom, []);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   983
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   984
          fun process_constr (dts, dt) (j, args1, args2) =
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   985
            let
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   986
              val Ts_idx =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   987
                map (Old_Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1);
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   988
              val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) Ts_idx;
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
   989
              val x =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   990
                Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt)
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
   991
                  (j + length dts);
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   992
            in
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
   993
              (j + length dts + 1,
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   994
               xs @ (x :: args1), fold_rev mk_abs_fun xs x :: args2)
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   995
            end;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   996
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   997
          val (_, args1, args2) = fold_rev process_constr dts (1, [], []);
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   998
          val Ts = map fastype_of args1;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
   999
          val c = list_comb (Const (cname, Ts ---> T), args1);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1000
          fun supp t =
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1001
            Const (@{const_name Nominal.supp}, fastype_of t --> HOLogic.mk_setT atomT) $ t;
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1002
          fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1003
          val supp_thm = Goal.prove_global_future thy8 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1004
            (augment_sort thy8 pt_cp_sort
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1005
              (HOLogic.mk_Trueprop (HOLogic.mk_eq
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1006
                (supp c,
30450
7655e6533209 HOLogic.mk_set, HOLogic.dest_set
haftmann
parents: 30364
diff changeset
  1007
                 if null dts then HOLogic.mk_set atomT []
35402
115a5a95710a clarified @{const_name} vs. @{const_abbrev};
wenzelm
parents: 35351
diff changeset
  1008
                 else foldr1 (HOLogic.mk_binop @{const_abbrev union}) (map supp args2)))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1009
            (fn {context = ctxt, ...} =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1010
              simp_tac (put_simpset HOL_basic_ss ctxt addsimps (supp_def ::
46161
4ed94d92ae19 prefer antiquotations;
wenzelm
parents: 45909
diff changeset
  1011
                 Un_assoc :: @{thm de_Morgan_conj} :: Collect_disj_eq :: finite_Un ::
45654
cf10bde35973 more antiquotations;
wenzelm
parents: 44693
diff changeset
  1012
                 Collect_False_empty :: finite_emptyI :: @{thms simp_thms} @
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1013
                 abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1014
        in
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1015
          (supp_thm,
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1016
           Goal.prove_global_future thy8 [] [] (augment_sort thy8 pt_cp_sort
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1017
             (HOLogic.mk_Trueprop (HOLogic.mk_eq
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1018
               (fresh c,
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45735
diff changeset
  1019
                if null dts then @{term True}
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1020
                else foldr1 HOLogic.mk_conj (map fresh args2)))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1021
             (fn {context = ctxt, ...} =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1022
               simp_tac (put_simpset HOL_ss ctxt addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1023
        end) atoms) constrs
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1024
      end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1025
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1026
    (**** weak induction theorem ****)
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1027
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1028
    fun mk_indrule_lemma (((i, _), T), U) (prems, concls) =
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1029
      let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1030
        val Rep_t = Const (nth rep_names i, T --> U) $ Old_Datatype_Aux.mk_Free "x" T i;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1031
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1032
        val Abs_t =  Const (nth abs_names i, U --> T);
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1033
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
  1034
      in
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
  1035
        (prems @ [HOLogic.imp $
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1036
            (Const (nth rep_set_names'' i, U --> HOLogic.boolT) $ Rep_t) $
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1037
              (Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
  1038
         concls @
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1039
           [Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1040
              Old_Datatype_Aux.mk_Free "x" T i])
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1041
      end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1042
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1043
    val (indrule_lemma_prems, indrule_lemma_concls) =
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1044
      fold mk_indrule_lemma (descr'' ~~ recTs ~~ recTs') ([], []);
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1045
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1046
    val indrule_lemma = Goal.prove_global_future thy8 [] []
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1047
      (Logic.mk_implies
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1048
        (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems),
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1049
         HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls)))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1050
         (fn {context = ctxt, ...} => EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1051
           [REPEAT (eresolve_tac ctxt [conjE] 1),
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1052
            REPEAT (EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1053
              [TRY (resolve_tac ctxt [conjI] 1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1054
               full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps Rep_inverse_thms) 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1055
               eresolve_tac ctxt [mp] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1056
               resolve_tac ctxt Rep_thms 1])]);
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1057
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1058
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule_lemma)));
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1059
    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
  1060
      map (Free o apfst fst o dest_Var) Ps;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1061
19833
3a3f591c838d - Changed naming scheme: names of "internal" constructors now have
berghofe
parents: 19710
diff changeset
  1062
    val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1063
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1064
    val dt_induct_prop = Old_Datatype_Prop.make_ind descr';
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1065
    val dt_induct = Goal.prove_global_future thy8 []
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1066
      (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
60781
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1067
      (fn {prems, context = ctxt} =>
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1068
        let
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1069
          val indrule_lemma' = infer_instantiate ctxt
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1070
            (map (#1 o dest_Var) Ps ~~ map (Thm.cterm_of ctxt) frees) indrule_lemma;
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1071
        in
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1072
          EVERY [resolve_tac ctxt [indrule_lemma'] 1,
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1073
           (Old_Datatype_Aux.ind_tac ctxt rep_induct [] THEN_ALL_NEW
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1074
             Object_Logic.atomize_prems_tac ctxt) 1,
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1075
           EVERY (map (fn (prem, r) => (EVERY
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1076
             [REPEAT (eresolve_tac ctxt Abs_inverse_thms' 1),
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1077
              simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Thm.symmetric r]) 1,
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1078
              DEPTH_SOLVE_1
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1079
                (assume_tac ctxt 1 ORELSE resolve_tac ctxt [prem] 1 ORELSE eresolve_tac ctxt [allE] 1)]))
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1080
                  (prems ~~ constr_defs))]
2da59cdf531c updated to infer_instantiate;
wenzelm
parents: 60754
diff changeset
  1081
        end);
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1082
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1083
    val case_names_induct = Old_Datatype_Data.mk_case_names_induct descr'';
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1084
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1085
    (**** prove that new datatypes have finite support ****)
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1086
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1087
    val _ = warning "proving finite support for the new datatype";
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents: 18245
diff changeset
  1088
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1089
    val indnames = Old_Datatype_Prop.make_tnames recTs;
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1090
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1091
    val abs_supp = Global_Theory.get_thms thy8 "abs_supp";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1092
    val supp_atm = Global_Theory.get_thms thy8 "supp_atm";
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1093
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1094
    val finite_supp_thms = map (fn atom =>
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1095
      let val atomT = Type (atom, [])
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
  1096
      in map Drule.export_without_context (List.take
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1097
        (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy8 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1098
           (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1099
             (HOLogic.mk_Trueprop
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1100
               (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1101
                 Const (@{const_name finite}, HOLogic.mk_setT atomT --> HOLogic.boolT) $
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1102
                   (Const (@{const_name Nominal.supp}, T --> HOLogic.mk_setT atomT) $ Free (s, T)))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1103
                   (indnames ~~ recTs)))))
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
  1104
           (fn {context = ctxt, ...} =>
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
  1105
             Old_Datatype_Aux.ind_tac ctxt dt_induct indnames 1 THEN
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
  1106
             ALLGOALS (asm_full_simp_tac (ctxt addsimps
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
  1107
              (abs_supp @ supp_atm @
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1108
               Global_Theory.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1109
               flat supp_thms))))),
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1110
         length new_type_names))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1111
      end) atoms;
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1112
18759
2f55e3e47355 Updated to Isabelle 2006-01-23
krauss
parents: 18707
diff changeset
  1113
    val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1114
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32957
diff changeset
  1115
        (* Function to add both the simp and eqvt attributes *)
22245
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1116
        (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1117
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1118
    val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  1119
 
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1120
    val (_, thy9) = thy8 |>
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24459
diff changeset
  1121
      Sign.add_path big_name |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1122
      Global_Theory.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
60359
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
  1123
      Global_Theory.add_thmss
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
  1124
        [((Binding.name "inducts", projections (Proof_Context.init_global thy8) dt_induct),
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
  1125
            [case_names_induct])] ||>
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24459
diff changeset
  1126
      Sign.parent_path ||>>
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1127
      Old_Datatype_Aux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1128
      Old_Datatype_Aux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1129
      Old_Datatype_Aux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1130
      Old_Datatype_Aux.store_thmss "inject" new_type_names inject_thms ||>>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1131
      Old_Datatype_Aux.store_thmss "supp" new_type_names supp_thms ||>>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1132
      Old_Datatype_Aux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1133
      fold (fn (atom, ths) => fn thy =>
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
  1134
        let
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
  1135
          val class = fs_class_of thy atom;
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
  1136
          val sort = Sign.minimize_sort thy (Sign.certify_sort thy (class :: pt_cp_sort));
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51671
diff changeset
  1137
        in fold (fn Type (s, Ts) => Axclass.prove_arity
28736
b1fd60fee652 Some modifications in code for proving arities to make it work for datatype
berghofe
parents: 28731
diff changeset
  1138
          (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1139
          (fn ctxt => Class.intro_classes_tac ctxt [] THEN resolve_tac ctxt ths 1)) newTs thy
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
  1140
        end) (atoms ~~ finite_supp_thms);
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1141
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1142
    (**** strong induction theorem ****)
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1143
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1144
    val pnames = if length descr'' = 1 then ["P"]
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1145
      else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1146
    val ind_sort = if null dt_atomTs then @{sort type}
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
  1147
      else Sign.minimize_sort thy9 (Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms));
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1148
    val fsT = TFree ("'n", ind_sort);
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1149
    val fsT' = TFree ("'n", @{sort type});
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1150
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1151
    val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1152
      (Old_Datatype_Prop.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1153
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1154
    fun make_pred fsT i T = Free (nth pnames i, fsT --> T --> HOLogic.boolT);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1155
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1156
    fun mk_fresh1 xs [] = []
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1157
      | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1158
            (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1159
              (filter (fn (_, U) => T = U) (rev xs)) @
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1160
          mk_fresh1 (y :: xs) ys;
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1161
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1162
    fun mk_fresh2 xss [] = []
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1163
      | mk_fresh2 xss ((p as (ys, _)) :: yss) = maps (fn y as (_, T) =>
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1164
            map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1165
              (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys @
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1166
          mk_fresh2 (p :: xss) yss;
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1167
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1168
    fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1169
      let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1170
        val recs = filter Old_Datatype_Aux.is_rec_type cargs;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1171
        val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') cargs;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1172
        val recTs' = map (Old_Datatype_Aux.typ_of_dtyp descr'') recs;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1173
        val tnames = Name.variant_list pnames (Old_Datatype_Prop.make_tnames Ts);
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1174
        val rec_tnames = map fst (filter (Old_Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs));
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1175
        val frees = tnames ~~ Ts;
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1176
        val frees' = partition_cargs idxs frees;
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42388
diff changeset
  1177
        val z = (singleton (Name.variant_list tnames) "z", fsT);
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1178
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1179
        fun mk_prem ((dt, s), T) =
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1180
          let
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1181
            val (Us, U) = strip_type T;
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
  1182
            val l = length Us;
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
  1183
          in
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 46215
diff changeset
  1184
            Logic.list_all (z :: map (pair "x") Us,
45838
653c84d5c6c9 do not open ML structures;
wenzelm
parents: 45822
diff changeset
  1185
              HOLogic.mk_Trueprop
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1186
                (make_pred fsT (Old_Datatype_Aux.body_index dt) U $ Bound l $
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1187
                  Old_Datatype_Aux.app_bnds (Free (s, T)) l))
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1188
          end;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1189
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1190
        val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1191
        val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1192
            (f T (Free p) (Free z))) (maps fst frees') @
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1193
          mk_fresh1 [] (maps fst frees') @
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1194
          mk_fresh2 [] frees'
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1195
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1196
      in
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1197
        fold_rev (Logic.all o Free) (frees @ [z])
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1198
          (Logic.list_implies (prems' @ prems,
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1199
            HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1200
              list_comb (Const (cname, Ts ---> T), map Free frees))))
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1201
      end;
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1202
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1203
    val ind_prems = maps (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1204
      map (make_ind_prem fsT (fn T => fn t => fn u =>
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1205
        fresh_const T fsT $ t $ u) i T)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1206
          (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs);
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1207
    val tnames = Old_Datatype_Prop.make_tnames recTs;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
  1208
    val zs = Name.variant_list tnames (replicate (length descr'') "z");
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38715
diff changeset
  1209
    val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1210
      (map (fn ((((i, _), T), tname), z) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1211
        make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1212
        (descr'' ~~ recTs ~~ tnames ~~ zs)));
18107
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1213
    val induct = Logic.list_implies (ind_prems, ind_concl);
ee6b4d3af498 Added strong induction theorem (currently only axiomatized!).
berghofe
parents: 18104
diff changeset
  1214
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1215
    val ind_prems' =
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1216
      map (fn (_, f as Free (_, T)) => Logic.all (Free ("x", fsT'))
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1217
        (HOLogic.mk_Trueprop (Const (@{const_name finite},
44692
ccfc7c193d2b modify nominal packages to better respect set/pred distinction
huffman
parents: 44241
diff changeset
  1218
          Term.range_type T -->
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
  1219
            HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1220
      maps (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1221
        map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1222
          HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1223
            (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs);
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38715
diff changeset
  1224
    val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1225
      (map (fn ((((i, _), T), tname), z) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1226
        make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1227
        (descr'' ~~ recTs ~~ tnames ~~ zs)));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1228
    val induct' = Logic.list_implies (ind_prems', ind_concl');
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1229
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1230
    val aux_ind_vars =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1231
      (Old_Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1232
       map mk_permT dt_atomTs) @ [("z", fsT')];
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1233
    val aux_ind_Ts = rev (map snd aux_ind_vars);
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38715
diff changeset
  1234
    val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1235
      (map (fn (((i, _), T), tname) =>
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1236
        HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1237
          fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1238
            (Free (tname, T))))
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1239
        (descr'' ~~ recTs ~~ tnames)));
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1240
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1241
    val fin_set_supp = map (fn s =>
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1242
      at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1243
    val fin_set_fresh = map (fn s =>
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1244
      at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1245
    val pt1_atoms = map (fn Type (s, _) =>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1246
      Global_Theory.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs;
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1247
    val pt2_atoms = map (fn Type (s, _) =>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1248
      Global_Theory.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs;
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1249
    val exists_fresh' = Global_Theory.get_thms thy9 "exists_fresh'";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1250
    val fs_atoms = Global_Theory.get_thms thy9 "fin_supp";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1251
    val abs_supp = Global_Theory.get_thms thy9 "abs_supp";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1252
    val perm_fresh_fresh = Global_Theory.get_thms thy9 "perm_fresh_fresh";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1253
    val calc_atm = Global_Theory.get_thms thy9 "calc_atm";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1254
    val fresh_atm = Global_Theory.get_thms thy9 "fresh_atm";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1255
    val fresh_left = Global_Theory.get_thms thy9 "fresh_left";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1256
    val perm_swap = Global_Theory.get_thms thy9 "perm_swap";
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1257
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1258
    fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1259
      let
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1260
        val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1261
        val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1262
            (HOLogic.exists_const T $ Abs ("x", T,
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1263
              fresh_const T (fastype_of p) $
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1264
                Bound 0 $ p)))
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1265
          (fn _ => EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1266
            [resolve_tac ctxt exists_fresh' 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1267
             simp_tac (put_simpset HOL_ss ctxt addsimps (supp_prod :: finite_Un :: fs_atoms @
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1268
               fin_set_supp @ ths)) 1]);
32202
443d5cfaba1b SUBPROOF/Obtain.result: named params;
wenzelm
parents: 32172
diff changeset
  1269
        val (([(_, cx)], ths), ctxt') = Obtain.result
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1270
          (fn ctxt' => EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1271
            [eresolve_tac ctxt' [exE] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1272
             full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1273
             REPEAT (eresolve_tac ctxt' [conjE] 1)])
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1274
          [ex] ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1275
      in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1276
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1277
    fun fresh_fresh_inst thy a b =
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1278
      let
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1279
        val T = fastype_of a;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1280
        val SOME th = find_first (fn th =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1281
          (case Thm.prop_of th of
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1282
            _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1283
          | _ => false)) perm_fresh_fresh
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1284
      in
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1285
        Drule.instantiate' []
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  1286
          [SOME (Thm.global_cterm_of thy a), NONE, SOME (Thm.global_cterm_of thy b)] th
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1287
      end;
18104
dbe58b104cb9 added thms perm, distinct and fresh to the simplifier.
urbanc
parents: 18068
diff changeset
  1288
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1289
    val fs_cp_sort =
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1290
      map (fs_class_of thy9) dt_atoms @
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33035
diff changeset
  1291
      maps (fn s => map (cp_class_of thy9 s) (remove (op =) s dt_atoms)) dt_atoms;
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1292
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1293
    (**********************************************************************
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1294
      The subgoals occurring in the proof of induct_aux have the
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1295
      following parameters:
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1296
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1297
        x_1 ... x_k p_1 ... p_m z
19710
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1298
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1299
      where
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1300
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1301
        x_i : constructor arguments (introduced by weak induction rule)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1302
        p_i : permutations (one for each atom type in the data type)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1303
        z   : freshness context
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1304
    ***********************************************************************)
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1305
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1306
    val _ = warning "proving strong induction theorem ...";
ee9c7fa80d21 Extended strong induction rule with additional
berghofe
parents: 19649
diff changeset
  1307
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1308
    val induct_aux = Goal.prove_global_future thy9 []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1309
        (map (augment_sort thy9 fs_cp_sort) ind_prems')
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1310
        (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1311
      let
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1312
        val (prems1, prems2) = chop (length dt_atomTs) prems;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1313
        val ind_ss2 = put_simpset HOL_ss context addsimps
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1314
          finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1315
        val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1316
          fresh_atm @ rev_simps @ app_simps;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1317
        val ind_ss3 = put_simpset HOL_ss context addsimps abs_fun_eq1 ::
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1318
          abs_perm @ calc_atm @ perm_swap;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1319
        val ind_ss4 = put_simpset HOL_basic_ss context addsimps fresh_left @ prems1 @
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1320
          fin_set_fresh @ calc_atm;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1321
        val ind_ss5 = put_simpset HOL_basic_ss context addsimps pt1_atoms;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1322
        val ind_ss6 = put_simpset HOL_basic_ss context addsimps flat perm_simps';
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1323
        val th = Goal.prove context [] []
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1324
          (augment_sort thy9 fs_cp_sort aux_ind_concl)
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1325
          (fn {context = context1, ...} =>
60328
9c94e6a30d29 clarified context;
wenzelm
parents: 60003
diff changeset
  1326
             EVERY (Old_Datatype_Aux.ind_tac context1 dt_induct tnames 1 ::
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1327
               maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1328
                 map (fn ((cname, cargs), is) =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1329
                   REPEAT (resolve_tac context1 [allI] 1) THEN
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1330
                   SUBPROOF (fn {prems = iprems, params, concl,
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1331
                       context = context2, ...} =>
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1332
                     let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1333
                       val concl' = Thm.term_of concl;
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1334
                       val _ $ (_ $ _ $ u) = concl';
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1335
                       val U = fastype_of u;
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1336
                       val (xs, params') =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1337
                         chop (length cargs) (map (Thm.term_of o #2) params);
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1338
                       val Ts = map fastype_of xs;
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1339
                       val cnstr = Const (cname, Ts ---> U);
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1340
                       val (pis, z) = split_last params';
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1341
                       val mk_pi = fold_rev (mk_perm []) pis;
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1342
                       val xs' = partition_cargs is xs;
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1343
                       val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs';
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1344
                       val ts = maps (fn (ts, u) => ts @ [u]) xs'';
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1345
                       val (freshs1, freshs2, context3) = fold (fn t =>
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1346
                         let val T = fastype_of t
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1347
                         in obtain_fresh_name' prems1
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1348
                           (the (AList.lookup op = fresh_fs T) $ z :: ts) T
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1349
                         end) (maps fst xs') ([], [], context2);
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1350
                       val freshs1' = unflat (map fst xs') freshs1;
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1351
                       val freshs2' = map (Simplifier.simplify ind_ss4)
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1352
                         (mk_not_sym freshs2);
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1353
                       val ind_ss1' = ind_ss1 addsimps freshs2';
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1354
                       val ind_ss3' = ind_ss3 addsimps freshs2';
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1355
                       val rename_eq =
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1356
                         if forall (null o fst) xs' then []
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1357
                         else [Goal.prove context3 [] []
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1358
                           (HOLogic.mk_Trueprop (HOLogic.mk_eq
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1359
                             (list_comb (cnstr, ts),
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1360
                              list_comb (cnstr, maps (fn ((bs, t), cs) =>
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1361
                                cs @ [fold_rev (mk_perm []) (map perm_of_pair
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1362
                                  (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1363
                           (fn {context = ctxt, ...} => EVERY
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1364
                              (simp_tac (put_simpset HOL_ss context3 addsimps flat inject_thms) 1 ::
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1365
                               REPEAT (FIRSTGOAL (resolve_tac ctxt [conjI])) ::
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1366
                               maps (fn ((bs, t), cs) =>
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1367
                                 if null bs then []
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1368
                                 else resolve_tac ctxt [sym] 1 :: maps (fn (b, c) =>
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1369
                                   [resolve_tac ctxt [trans] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1370
                                    resolve_tac ctxt [sym] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1371
                                    resolve_tac ctxt [fresh_fresh_inst thy9 b c] 1,
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1372
                                    simp_tac ind_ss1' 1,
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1373
                                    simp_tac ind_ss2 1,
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1374
                                    simp_tac ind_ss3' 1]) (bs ~~ cs))
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1375
                                 (xs'' ~~ freshs1')))];
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1376
                       val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1377
                         [simp_tac (ind_ss6 addsimps rename_eq) 1,
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1378
                          cut_facts_tac iprems 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1379
                          (resolve_tac context2 prems THEN_ALL_NEW
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1380
                            SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1381
                                _ $ (Const (@{const_name Nominal.fresh}, _) $ _ $ _) =>
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1382
                                  simp_tac ind_ss1' i
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  1383
                              | _ $ (Const (@{const_name Not}, _) $ _) =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1384
                                  resolve_tac context2 freshs2' i
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1385
                              | _ => asm_simp_tac (put_simpset HOL_basic_ss context3 addsimps
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1386
                                  pt2_atoms addsimprocs [perm_simproc]) i)) 1])
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41489
diff changeset
  1387
                       val final = Proof_Context.export context3 context2 [th]
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1388
                     in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1389
                       resolve_tac context2 final 1
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1390
                     end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1391
      in
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1392
        EVERY
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1393
          [cut_facts_tac [th] 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1394
           REPEAT (eresolve_tac context [conjE, @{thm allE_Nil}] 1),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1395
           REPEAT (eresolve_tac context [allE] 1),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1396
           REPEAT (TRY (resolve_tac context [conjI] 1) THEN asm_full_simp_tac ind_ss5 1)]
25951
6ebe26bfed18 Reimplemented proof of strong induction theorem.
berghofe
parents: 25823
diff changeset
  1397
      end);
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1398
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1399
    val induct_aux' = Thm.instantiate ([],
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60359
diff changeset
  1400
      map (fn (s, Var (v as (_, T))) =>
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60359
diff changeset
  1401
        (v, Thm.global_cterm_of thy9 (Free (s, T))))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1402
          (pnames ~~ map head_of (HOLogic.dest_conj
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1403
             (HOLogic.dest_Trueprop (Thm.concl_of induct_aux)))) @
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1404
      map (fn (_, f) =>
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35842
diff changeset
  1405
        let val f' = Logic.varify_global f
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60359
diff changeset
  1406
        in (dest_Var f',
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  1407
          Thm.global_cterm_of thy9 (Const (@{const_name Nominal.supp}, fastype_of f')))
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1408
        end) fresh_fs) induct_aux;
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1409
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1410
    val induct = Goal.prove_global_future thy9 []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1411
      (map (augment_sort thy9 fs_cp_sort) ind_prems)
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1412
      (augment_sort thy9 fs_cp_sort ind_concl)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1413
      (fn {prems, context = ctxt} => EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1414
         [resolve_tac ctxt [induct_aux'] 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1415
          REPEAT (resolve_tac ctxt fs_atoms 1),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1416
          REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1417
            (eresolve_tac ctxt @{thms meta_spec} ORELSE'
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1418
              full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [fresh_def]))) 1)])
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1419
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1420
    val (_, thy10) = thy9 |>
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24459
diff changeset
  1421
      Sign.add_path big_name |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1422
      Global_Theory.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1423
      Global_Theory.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
60359
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
  1424
      Global_Theory.add_thmss
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
  1425
        [((Binding.name "strong_inducts", projections (Proof_Context.init_global thy9) induct),
cb8828b859a1 clarified context;
wenzelm
parents: 60328
diff changeset
  1426
            [case_names_induct])];
18658
317a6f0ef8b9 Implemented proof of (strong) induction rule.
berghofe
parents: 18582
diff changeset
  1427
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1428
    (**** recursion combinator ****)
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1429
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1430
    val _ = warning "defining recursion combinator ...";
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1431
45741
088256c289e7 eliminated some legacy operations;
wenzelm
parents: 45740
diff changeset
  1432
    val used = fold Term.add_tfree_namesT recTs [];
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1433
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1434
    val (rec_result_Ts', rec_fn_Ts') = Old_Datatype_Prop.make_primrec_Ts descr' used;
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1435
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1436
    val rec_sort = if null dt_atomTs then @{sort type} else
36428
874843c1e96e really minimize sorts after certification -- looks like this is intended here;
wenzelm
parents: 35994
diff changeset
  1437
      Sign.minimize_sort thy10 (Sign.certify_sort thy10 pt_cp_sort);
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1438
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1439
    val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1440
    val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1441
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1442
    val rec_set_Ts = map (fn (T1, T2) =>
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1443
      rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1444
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1445
    val big_rec_name = big_name ^ "_rec_set";
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1446
    val rec_set_names' =
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1447
      if length descr'' = 1 then [big_rec_name] else
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1448
        map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1449
          (1 upto (length descr''));
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
  1450
    val rec_set_names =  map (Sign.full_bname thy10) rec_set_names';
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1451
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1452
    val rec_fns = map (uncurry (Old_Datatype_Aux.mk_Free "f"))
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1453
      (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1454
    val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1455
      (rec_set_names' ~~ rec_set_Ts);
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1456
    val rec_sets = map (fn c => list_comb (Const c, rec_fns))
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1457
      (rec_set_names ~~ rec_set_Ts);
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1458
19322
bf84bdf05f14 Replaced iteration combinator by recursion combinator.
berghofe
parents: 19275
diff changeset
  1459
    (* introduction rules for graph of recursion function *)
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1460
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1461
    val rec_preds = map (fn (a, T) =>
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1462
      Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1463
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1464
    fun mk_fresh3 rs [] = []
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1465
      | mk_fresh3 rs ((p as (ys, z)) :: yss) = maps (fn y as (_, T) =>
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1466
            map_filter (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1467
              else SOME (HOLogic.mk_Trueprop
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1468
                (fresh_const T U $ Free y $ Free r))) rs) ys @
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1469
          mk_fresh3 rs yss;
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1470
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1471
    (* FIXME: avoid collisions with other variable names? *)
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1472
    val rec_ctxt = Free ("z", fsT');
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1473
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1474
    fun make_rec_intr T p rec_set ((cname, cargs), idxs)
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1475
        (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, l) =
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1476
      let
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1477
        val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') cargs;
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1478
        val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1479
        val frees' = partition_cargs idxs frees;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1480
        val binders = maps fst frees';
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1481
        val atomTs = distinct op = (maps (map snd o fst) frees');
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1482
        val recs = map_filter
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1483
          (fn ((_, Old_Datatype_Aux.DtRec i), p) => SOME (i, p) | _ => NONE)
19851
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1484
          (partition_cargs idxs cargs ~~ frees');
10162c01bd78 Completely rewrote code for defining graph of recursion combinator.
berghofe
parents: 19833
diff changeset
  1485
        val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1486
          map (fn (i, _) => nth rec_result_Ts i) recs;
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1487
        val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1488
          (nth rec_sets' i $ Free x $ Free y)) (recs ~~ frees'');
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1489
        val prems2 =
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1490
          map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1491
            (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1492
        val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1493
        val prems4 = map (fn ((i, _), y) =>
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1494
          HOLogic.mk_Trueprop (nth rec_preds i $ Free y)) (recs ~~ frees'');
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1495
        val prems5 = mk_fresh3 (recs ~~ frees'') frees';
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1496
        val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1497
          (Const (@{const_name finite}, HOLogic.mk_setT aT --> HOLogic.boolT) $
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1498
             (Const (@{const_name Nominal.supp}, T --> HOLogic.mk_setT aT) $ Free y)))
20411
dd8a1cda2eaf Added premises concerning finite support of recursion results to FCBs.
berghofe
parents: 20397
diff changeset
  1499
               frees'') atomTs;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1500
        val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1501
          (fresh_const T fsT' $ Free x $ rec_ctxt)) binders;
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1502
        val result = list_comb (nth rec_fns l, map Free (frees @ frees''));
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1503
        val result_freshs = map (fn p as (_, T) =>
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1504
          fresh_const T (fastype_of result) $ Free p $ result) binders;
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1505
        val P = HOLogic.mk_Trueprop (p $ result)
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1506
      in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1507
        (rec_intr_ts @ [Logic.list_implies (flat prems2 @ prems3 @ prems1,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1508
           HOLogic.mk_Trueprop (rec_set $
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1509
             list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1510
         rec_prems @ [fold_rev (Logic.all o Free) (frees @ frees'') (Logic.list_implies (prems4, P))],
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1511
         rec_prems' @ map (fn fr => fold_rev (Logic.all o Free) (frees @ frees'')
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 46161
diff changeset
  1512
           (Logic.list_implies (nth prems2 l @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
21054
6048c085c3ae Split up FCBs into separate formulae for each binder.
berghofe
parents: 21021
diff changeset
  1513
             HOLogic.mk_Trueprop fr))) result_freshs,
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1514
         rec_eq_prems @ [flat prems2 @ prems3],
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1515
         l + 1)
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1516
      end;
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1517
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1518
    val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1519
      fold (fn ((((d, d'), T), p), rec_set) =>
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1520
        fold (make_rec_intr T p rec_set) (#3 (snd d) ~~ snd d'))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1521
          (descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets')
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
  1522
          ([], [], [], [], 0);
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1523
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  1524
    val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
  1525
      thy10
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
  1526
      |> Sign.concealed
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33669
diff changeset
  1527
      |> Inductive.add_inductive_global
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
  1528
          {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 46961
diff changeset
  1529
           coind = false, no_elim = false, no_ind = false, skip_mono = true}
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
  1530
          (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
26128
fe2d24c26e0c inductive package: simplified group handling;
wenzelm
parents: 26067
diff changeset
  1531
          (map dest_Free rec_fns)
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
  1532
          (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) []
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1533
      ||> Global_Theory.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct")
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
  1534
      ||> Sign.restore_naming thy10;
19251
6bc0dda66f32 First version of function for defining graph of iteration combinator.
berghofe
parents: 19134
diff changeset
  1535
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1536
    (** equivariance **)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1537
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1538
    val fresh_bij = Global_Theory.get_thms thy11 "fresh_bij";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1539
    val perm_bij = Global_Theory.get_thms thy11 "perm_bij";
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1540
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1541
    val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1542
      let
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1543
        val permT = mk_permT aT;
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1544
        val pi = Free ("pi", permT);
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1545
        val rec_fns_pi = map (mk_perm [] pi o uncurry (Old_Datatype_Aux.mk_Free "f"))
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1546
          (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1547
        val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1548
          (rec_set_names ~~ rec_set_Ts);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1549
        val ps = map (fn ((((T, U), R), R'), i) =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1550
          let
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1551
            val x = Free ("x" ^ string_of_int i, T);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1552
            val y = Free ("y" ^ string_of_int i, U)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1553
          in
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1554
            (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1555
          end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1556
        val ths = map (fn th => Drule.export_without_context (th RS mp))
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1557
          (Old_Datatype_Aux.split_conj_thm
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1558
            (Goal.prove_global_future thy11 [] []
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1559
              (augment_sort thy1 pt_cp_sort
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1560
                (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1561
              (fn {context = ctxt, ...} =>
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1562
                 resolve_tac ctxt [rec_induct] 1 THEN REPEAT
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1563
                 (simp_tac (put_simpset HOL_basic_ss ctxt
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1564
                    addsimps flat perm_simps'
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1565
                    addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1566
                  (resolve_tac ctxt rec_intrs THEN_ALL_NEW
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1567
                   asm_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_bij @ perm_bij))) 1))))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19985
diff changeset
  1568
        val ths' = map (fn ((P, Q), th) =>
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1569
          Goal.prove_global_future thy11 [] []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1570
            (augment_sort thy1 pt_cp_sort
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1571
              (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1572
            (fn {context = ctxt, ...} =>
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1573
              dresolve_tac ctxt [Thm.instantiate ([],
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60359
diff changeset
  1574
                 [((("pi", 0), permT),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1575
                   Thm.global_cterm_of thy11 (Const (@{const_name rev}, permT --> permT) $ pi))]) th] 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1576
               NominalPermeq.perm_simp_tac (put_simpset HOL_ss ctxt) 1)) (ps ~~ ths)
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1577
      in (ths, ths') end) dt_atomTs);
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1578
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1579
    (** finite support **)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1580
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1581
    val rec_fin_supp_thms = map (fn aT =>
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1582
      let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
  1583
        val name = Long_Name.base_name (fst (dest_Type aT));
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1584
        val fs_name = Global_Theory.get_thm thy11 ("fs_" ^ name ^ "1");
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1585
        val aset = HOLogic.mk_setT aT;
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1586
        val finite = Const (@{const_name finite}, aset --> HOLogic.boolT);
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1587
        val fins = map (fn (f, T) => HOLogic.mk_Trueprop
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1588
          (finite $ (Const (@{const_name Nominal.supp}, T --> aset) $ f)))
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1589
            (rec_fns ~~ rec_fn_Ts)
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1590
      in
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1591
        map (fn th => Drule.export_without_context (th RS mp)) (Old_Datatype_Aux.split_conj_thm
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  1592
          (Goal.prove_global_future thy11 []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1593
            (map (augment_sort thy11 fs_cp_sort) fins)
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1594
            (augment_sort thy11 fs_cp_sort
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1595
              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1596
                (map (fn (((T, U), R), i) =>
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1597
                   let
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1598
                     val x = Free ("x" ^ string_of_int i, T);
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1599
                     val y = Free ("y" ^ string_of_int i, U)
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1600
                   in
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1601
                     HOLogic.mk_imp (R $ x $ y,
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1602
                       finite $ (Const (@{const_name Nominal.supp}, U --> aset) $ y))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1603
                   end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1604
                     (1 upto length recTs))))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1605
            (fn {prems = fins, context = ctxt} =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1606
              (resolve_tac ctxt [rec_induct] THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1607
               (NominalPermeq.finite_guess_tac (put_simpset HOL_ss ctxt addsimps [fs_name]) 1))))
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1608
      end) dt_atomTs;
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  1609
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1610
    (** freshness **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1611
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1612
    val finite_premss = map (fn aT =>
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1613
      map (fn (f, T) => HOLogic.mk_Trueprop
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1614
        (Const (@{const_name finite}, HOLogic.mk_setT aT --> HOLogic.boolT) $
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1615
           (Const (@{const_name Nominal.supp}, T --> HOLogic.mk_setT aT) $ f)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1616
           (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1617
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1618
    val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1619
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1620
    val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1621
      let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
  1622
        val name = Long_Name.base_name (fst (dest_Type aT));
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  1623
        val fs_name = Global_Theory.get_thm thy11 ("fs_" ^ name ^ "1");
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1624
        val a = Free ("a", aT);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1625
        val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1626
          (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1627
      in
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1628
        map (fn (((T, U), R), eqvt_th) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1629
          let
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1630
            val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1631
            val y = Free ("y", U);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1632
            val y' = Free ("y'", U)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1633
          in
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41489
diff changeset
  1634
            Drule.export_without_context (Goal.prove (Proof_Context.init_global thy11) []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1635
              (map (augment_sort thy11 fs_cp_sort)
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1636
                (finite_prems @
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1637
                   [HOLogic.mk_Trueprop (R $ x $ y),
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1638
                    HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1639
                      HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1640
                    HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1641
                 freshs))
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1642
              (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1643
              (fn {prems, context} =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1644
                 let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1645
                   val (finite_prems, rec_prem :: unique_prem ::
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1646
                     fresh_prems) = chop (length finite_prems) prems;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1647
                   val unique_prem' = unique_prem RS spec RS mp;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1648
                   val unique = [unique_prem', unique_prem' RS sym] MRS trans;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1649
                   val _ $ (_ $ (_ $ S $ _)) $ _ = Thm.prop_of supports_fresh;
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1650
                   val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1651
                 in EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1652
                   [resolve_tac context [Drule.cterm_instantiate
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  1653
                      [(Thm.global_cterm_of thy11 S,
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  1654
                        Thm.global_cterm_of thy11 (Const (@{const_name Nominal.supp},
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1655
                          fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1656
                      supports_fresh] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1657
                    simp_tac (put_simpset HOL_basic_ss context addsimps
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
  1658
                      [supports_def, Thm.symmetric fresh_def, fresh_prod]) 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1659
                    REPEAT_DETERM (resolve_tac context [allI, impI] 1),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1660
                    REPEAT_DETERM (eresolve_tac context [conjE] 1),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1661
                    resolve_tac context [unique] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1662
                    SUBPROOF (fn {context = ctxt, prems = prems', params = [(_, a), (_, b)], ...} =>
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1663
                      EVERY [cut_facts_tac [rec_prem] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1664
                       resolve_tac ctxt [Thm.instantiate ([],
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60359
diff changeset
  1665
                         [((("pi", 0), mk_permT aT),
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  1666
                           Thm.global_cterm_of thy11
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1667
                            (perm_of_pair (Thm.term_of a, Thm.term_of b)))]) eqvt_th] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1668
                       asm_simp_tac (put_simpset HOL_ss context addsimps
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1669
                         (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1670
                    resolve_tac context [rec_prem] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1671
                    simp_tac (put_simpset HOL_ss context addsimps (fs_name ::
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1672
                      supp_prod :: finite_Un :: finite_prems)) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1673
                    simp_tac (put_simpset HOL_ss context addsimps (Thm.symmetric fresh_def ::
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1674
                      fresh_prod :: fresh_prems)) 1]
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1675
                 end))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1676
          end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1677
      end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1678
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1679
    (** uniqueness **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1680
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1681
    val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1682
    val fun_tupleT = fastype_of fun_tuple;
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1683
    val rec_unique_frees =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1684
      Old_Datatype_Prop.indexify_names (replicate (length recTs) "x") ~~ recTs;
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1685
    val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1686
    val rec_unique_frees' =
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1687
      Old_Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1688
    val rec_unique_concls = map (fn ((x, U), R) =>
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  1689
        Const (@{const_name Ex1}, (U --> HOLogic.boolT) --> HOLogic.boolT) $
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1690
          Abs ("y", U, R $ Free x $ Bound 0))
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1691
      (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1692
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1693
    val induct_aux_rec = Drule.cterm_instantiate
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  1694
      (map (apply2 (Thm.global_cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35842
diff changeset
  1695
         (map (fn (aT, f) => (Logic.varify_global f, Abs ("z", HOLogic.unitT,
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1696
            Const (@{const_name Nominal.supp}, fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1697
              fresh_fs @
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1698
          map (fn (((P, T), (x, U)), Q) =>
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35842
diff changeset
  1699
           (Var ((P, 0), Logic.varifyT_global (fsT' --> T --> HOLogic.boolT)),
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
  1700
            Abs ("z", HOLogic.unitT, absfree (x, U) Q)))
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1701
              (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35842
diff changeset
  1702
          map (fn (s, T) => (Var ((s, 0), Logic.varifyT_global T), Free (s, T)))
20145
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1703
            rec_unique_frees)) induct_aux;
922b4e7b8efd Started implementing uniqueness proof for recursion
berghofe
parents: 20071
diff changeset
  1704
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1705
    fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1706
      let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1707
        val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1708
        val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1709
            (HOLogic.exists_const T $ Abs ("x", T,
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1710
              fresh_const T (fastype_of p) $ Bound 0 $ p)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1711
          (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1712
            [cut_facts_tac ths 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1713
             REPEAT_DETERM (dresolve_tac ctxt (the (AList.lookup op = rec_fin_supp T)) 1),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1714
             resolve_tac ctxt exists_fresh' 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1715
             asm_simp_tac (put_simpset HOL_ss ctxt addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
32202
443d5cfaba1b SUBPROOF/Obtain.result: named params;
wenzelm
parents: 32172
diff changeset
  1716
        val (([(_, cx)], ths), ctxt') = Obtain.result
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1717
          (fn ctxt' => EVERY
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1718
            [eresolve_tac ctxt' [exE] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1719
             full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1720
             REPEAT (eresolve_tac ctxt [conjE] 1)])
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1721
          [ex] ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1722
      in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1723
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1724
    val finite_ctxt_prems = map (fn aT =>
22274
ce1459004c8d Adapted to changes in Finite_Set theory.
berghofe
parents: 22245
diff changeset
  1725
      HOLogic.mk_Trueprop
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1726
        (Const (@{const_name finite}, HOLogic.mk_setT aT --> HOLogic.boolT) $
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1727
           (Const (@{const_name Nominal.supp}, fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1728
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1729
    val rec_unique_thms = Old_Datatype_Aux.split_conj_thm (Goal.prove
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41489
diff changeset
  1730
      (Proof_Context.init_global thy11) (map fst rec_unique_frees)
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1731
      (map (augment_sort thy11 fs_cp_sort)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1732
        (flat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1733
      (augment_sort thy11 fs_cp_sort
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1734
        (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1735
      (fn {prems, context} =>
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1736
         let
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1737
           val k = length rec_fns;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1738
           val (finite_thss, ths1) = fold_map (fn T => fn xs =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1739
             apfst (pair T) (chop k xs)) dt_atomTs prems;
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1740
           val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1741
           val (P_ind_ths, fcbs) = chop k ths2;
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  1742
           val P_ths = map (fn th => th RS mp) (Old_Datatype_Aux.split_conj_thm
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1743
             (Goal.prove context
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1744
               (map fst (rec_unique_frees'' @ rec_unique_frees')) []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1745
               (augment_sort thy11 fs_cp_sort
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1746
                 (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1747
                    (map (fn (((x, y), S), P) => HOLogic.mk_imp
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1748
                      (S $ Free x $ Free y, P $ (Free y)))
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1749
                        (rec_unique_frees'' ~~ rec_unique_frees' ~~
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1750
                           rec_sets ~~ rec_preds)))))
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58959
diff changeset
  1751
               (fn {context = ctxt, ...} =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1752
                  resolve_tac ctxt [rec_induct] 1 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1753
                  REPEAT ((resolve_tac ctxt P_ind_ths THEN_ALL_NEW assume_tac ctxt) 1))));
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1754
           val rec_fin_supp_thms' = map
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1755
             (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1756
             (rec_fin_supp_thms ~~ finite_thss);
20267
1154363129be Additional freshness constraints for FCB.
berghofe
parents: 20179
diff changeset
  1757
         in EVERY
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1758
           ([resolve_tac context [induct_aux_rec] 1] @
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1759
            maps (fn ((_, finite_ths), finite_th) =>
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1760
              [cut_facts_tac (finite_th :: finite_ths) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1761
               asm_simp_tac (put_simpset HOL_ss context addsimps [supp_prod, finite_Un]) 1])
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  1762
                (finite_thss ~~ finite_ctxt_ths) @
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1763
            maps (fn ((_, idxss), elim) => maps (fn idxs =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1764
              [full_simp_tac (put_simpset HOL_ss context addsimps [Thm.symmetric fresh_def, supp_prod, Un_iff]) 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1765
               REPEAT_DETERM (eresolve_tac context @{thms conjE ex1E} 1),
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1766
               resolve_tac context @{thms ex1I} 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1767
               (resolve_tac context rec_intrs THEN_ALL_NEW assume_tac context) 1,
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1768
               rotate_tac ~1 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1769
               ((DETERM o eresolve_tac context [elim]) THEN_ALL_NEW full_simp_tac
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1770
                  (put_simpset HOL_ss context addsimps flat distinct_thms)) 1] @
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
  1771
               (if null idxs then [] else [hyp_subst_tac context 1,
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  1772
                SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1773
                  let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1774
                    val SOME prem = find_first (can (HOLogic.dest_eq o
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1775
                      HOLogic.dest_Trueprop o Thm.prop_of)) prems';
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1776
                    val _ $ (_ $ lhs $ rhs) = Thm.prop_of prem;
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1777
                    val _ $ (_ $ lhs' $ rhs') = Thm.term_of concl;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1778
                    val rT = fastype_of lhs';
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1779
                    val (c, cargsl) = strip_comb lhs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1780
                    val cargsl' = partition_cargs idxs cargsl;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1781
                    val boundsl = maps fst cargsl';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1782
                    val (_, cargsr) = strip_comb rhs;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1783
                    val cargsr' = partition_cargs idxs cargsr;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1784
                    val boundsr = maps fst cargsr';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1785
                    val (params1, _ :: params2) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1786
                      chop (length params div 2) (map (Thm.term_of o #2) params);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1787
                    val params' = params1 @ params2;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1788
                    val rec_prems = filter (fn th =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1789
                        (case Thm.prop_of th of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1790
                          _ $ p =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1791
                            (case head_of p of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1792
                              Const (s, _) => member (op =) rec_set_names s
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1793
                            | _ => false)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1794
                        | _ => false)) prems';
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1795
                    val fresh_prems = filter (fn th =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1796
                      (case Thm.prop_of th of
56253
83b3c110f22d more antiquotations;
wenzelm
parents: 56239
diff changeset
  1797
                        _ $ (Const (@{const_name Nominal.fresh}, _) $ _ $ _) => true
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  1798
                      | _ $ (Const (@{const_name Not}, _) $ _) => true
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1799
                      | _ => false)) prems';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1800
                    val Ts = map fastype_of boundsl;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1801
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1802
                    val _ = warning "step 1: obtaining fresh names";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1803
                    val (freshs1, freshs2, context'') = fold
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1804
                      (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1805
                         (maps snd finite_thss @ finite_ctxt_ths @ rec_prems)
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1806
                         rec_fin_supp_thms')
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1807
                      Ts ([], [], context');
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1808
                    val pi1 = map perm_of_pair (boundsl ~~ freshs1);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1809
                    val rpi1 = rev pi1;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1810
                    val pi2 = map perm_of_pair (boundsr ~~ freshs1);
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1811
                    val rpi2 = rev pi2;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1812
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1813
                    val fresh_prems' = mk_not_sym fresh_prems;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1814
                    val freshs2' = mk_not_sym freshs2;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1815
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1816
                    (** as, bs, cs # K as ts, K bs us **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1817
                    val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1818
                    val prove_fresh_simpset = put_simpset HOL_ss context'' addsimps
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1819
                      (finite_Diff :: flat fresh_thms @
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1820
                       fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1821
                    (* FIXME: avoid asm_full_simp_tac ? *)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1822
                    fun prove_fresh ths y x = Goal.prove context'' [] []
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1823
                      (HOLogic.mk_Trueprop (fresh_const
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1824
                         (fastype_of x) (fastype_of y) $ x $ y))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1825
                      (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_simpset 1);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1826
                    val constr_fresh_thms =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1827
                      map (prove_fresh fresh_prems lhs) boundsl @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1828
                      map (prove_fresh fresh_prems rhs) boundsr @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1829
                      map (prove_fresh freshs2 lhs) freshs1 @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1830
                      map (prove_fresh freshs2 rhs) freshs1;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1831
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1832
                    (** pi1 o (K as ts) = pi2 o (K bs us) **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1833
                    val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1834
                    val pi1_pi2_eq = Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1835
                      (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1836
                        (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1837
                      (fn {context = ctxt, ...} => EVERY
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1838
                         [cut_facts_tac constr_fresh_thms 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1839
                          asm_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh) 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1840
                          resolve_tac ctxt [prem] 1]);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1841
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1842
                    (** pi1 o ts = pi2 o us **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1843
                    val _ = warning "step 4: pi1 o ts = pi2 o us";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1844
                    val pi1_pi2_eqs = map (fn (t, u) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1845
                      Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1846
                        (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1847
                          (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1848
                        (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1849
                           [cut_facts_tac [pi1_pi2_eq] 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1850
                            asm_full_simp_tac (put_simpset HOL_ss context'' addsimps
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1851
                              (calc_atm @ flat perm_simps' @
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1852
                               fresh_prems' @ freshs2' @ abs_perm @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1853
                               alpha @ flat inject_thms)) 1]))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1854
                        (map snd cargsl' ~~ map snd cargsr');
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1855
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1856
                    (** pi1^-1 o pi2 o us = ts **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1857
                    val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1858
                    val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1859
                      Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1860
                        (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1861
                          (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1862
                        (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1863
                           ((eq RS sym) :: perm_swap)) 1))
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1864
                        (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1865
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1866
                    val (rec_prems1, rec_prems2) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1867
                      chop (length rec_prems div 2) rec_prems;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1868
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1869
                    (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1870
                    val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1871
                    val rec_prems' = map (fn th =>
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1872
                      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1873
                        val _ $ (S $ x $ y) = Thm.prop_of th;
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1874
                        val Const (s, _) = head_of S;
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  1875
                        val k = find_index (equal s) rec_set_names;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1876
                        val pi = rpi1 @ pi2;
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1877
                        fun mk_pi z = fold_rev (mk_perm []) pi z;
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1878
                        fun eqvt_tac ctxt p =
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1879
                          let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1880
                            val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1881
                            val l = find_index (equal T) dt_atomTs;
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1882
                            val th = nth (nth rec_equiv_thms' l) k;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1883
                            val th' = Thm.instantiate ([],
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60359
diff changeset
  1884
                              [((("pi", 0), U), Thm.global_cterm_of thy11 p)]) th;
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1885
                          in resolve_tac ctxt [th'] 1 end;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1886
                        val th' = Goal.prove context'' [] []
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1887
                          (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1888
                          (fn {context = ctxt, ...} => EVERY
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1889
                             (map (eqvt_tac ctxt) pi @
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1890
                              [simp_tac (put_simpset HOL_ss context'' addsimps (fresh_prems' @ freshs2' @
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1891
                                 perm_swap @ perm_fresh_fresh)) 1,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1892
                               resolve_tac ctxt [th] 1]))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1893
                      in
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1894
                        Simplifier.simplify
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1895
                          (put_simpset HOL_basic_ss context'' addsimps rpi1_pi2_eqs) th'
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1896
                      end) rec_prems2;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1897
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1898
                    val ihs = filter (fn th =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1899
                      (case Thm.prop_of th of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1900
                        _ $ (Const (@{const_name All}, _) $ _) => true
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1901
                      | _ => false)) prems';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1902
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1903
                    (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1904
                    val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1905
                    val rec_eqns = map (fn (th, ih) =>
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1906
                      let
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1907
                        val th' = th RS (ih RS spec RS mp) RS sym;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1908
                        val _ $ (_ $ lhs $ rhs) = Thm.prop_of th';
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1909
                        fun strip_perm (_ $ _ $ t) = strip_perm t
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1910
                          | strip_perm t = t;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1911
                      in
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1912
                        Goal.prove context'' [] []
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1913
                           (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1914
                              (fold_rev (mk_perm []) pi1 lhs,
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1915
                               fold_rev (mk_perm []) pi2 (strip_perm rhs))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1916
                           (fn _ => simp_tac (put_simpset HOL_basic_ss context'' addsimps
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1917
                              (th' :: perm_swap)) 1)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1918
                      end) (rec_prems' ~~ ihs);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1919
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1920
                    (** as # rs **)
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1921
                    val _ = warning "step 8: as # rs";
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1922
                    val rec_freshs =
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1923
                      maps (fn (rec_prem, ih) =>
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1924
                        let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20548
diff changeset
  1925
                          val _ $ (S $ x $ (y as Free (_, T))) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  1926
                            Thm.prop_of rec_prem;
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1927
                          val k = find_index (equal S) rec_sets;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1928
                          val atoms = flat (map_filter (fn (bs, z) =>
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1929
                            if z = x then NONE else SOME bs) cargsl')
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1930
                        in
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1931
                          map (fn a as Free (_, aT) =>
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1932
                            let val l = find_index (equal aT) dt_atomTs;
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1933
                            in
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1934
                              Goal.prove context'' [] []
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1935
                                (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1936
                                (fn {context = ctxt, ...} => EVERY
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1937
                                   (resolve_tac ctxt [nth (nth rec_fresh_thms l) k] 1 ::
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1938
                                    map (fn th => resolve_tac ctxt [th] 1)
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
  1939
                                      (snd (nth finite_thss l)) @
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1940
                                    [resolve_tac ctxt [rec_prem] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1941
                                     resolve_tac ctxt [ih] 1,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1942
                                     REPEAT_DETERM (resolve_tac ctxt fresh_prems 1)]))
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1943
                            end) atoms
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1944
                        end) (rec_prems1 ~~ ihs);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1945
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1946
                    (** as # fK as ts rs , bs # fK bs us vs **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1947
                    val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1948
                    fun prove_fresh_result (a as Free (_, aT)) =
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1949
                      Goal.prove context'' [] []
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1950
                        (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1951
                        (fn _ => EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1952
                           [resolve_tac context'' fcbs 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1953
                            REPEAT_DETERM (resolve_tac context''
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1954
                              (fresh_prems @ rec_freshs) 1),
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1955
                            REPEAT_DETERM (resolve_tac context'' (maps snd rec_fin_supp_thms') 1
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1956
                              THEN resolve_tac context'' rec_prems 1),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1957
                            resolve_tac context'' P_ind_ths 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1958
                            REPEAT_DETERM (resolve_tac context'' (P_ths @ rec_prems) 1)]);
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  1959
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1960
                    val fresh_results'' = map prove_fresh_result boundsl;
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1961
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1962
                    fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1963
                      let val th' = Goal.prove context'' [] []
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1964
                        (HOLogic.mk_Trueprop (fresh_const aT rT $
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1965
                            fold_rev (mk_perm []) (rpi2 @ pi1) a $
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  1966
                            fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1967
                        (fn {context = ctxt, ...} =>
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1968
                          simp_tac (put_simpset HOL_ss context'' addsimps fresh_bij) 1 THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1969
                          resolve_tac ctxt [th] 1)
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1970
                      in
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1971
                        Goal.prove context'' [] []
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1972
                          (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
  1973
                          (fn {context = ctxt, ...} => EVERY
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1974
                             [cut_facts_tac [th'] 1,
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
  1975
                              full_simp_tac (put_simpset HOL_ss ctxt
25998
f38dc602a926 Tuned uniqueness proof for recursion combinator.
berghofe
parents: 25985
diff changeset
  1976
                                addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap
f38dc602a926 Tuned uniqueness proof for recursion combinator.
berghofe
parents: 25985
diff changeset
  1977
                                addsimprocs [NominalPermeq.perm_simproc_app]) 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1978
                              full_simp_tac (put_simpset HOL_ss context'' addsimps (calc_atm @
21073
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1979
                                fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1980
                      end;
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1981
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1982
                    val fresh_results = fresh_results'' @ map prove_fresh_result''
be0a17371ba6 Proof of "bs # fK bs us vs" no longer depends on FCBs.
berghofe
parents: 21055
diff changeset
  1983
                      (boundsl ~~ boundsr ~~ fresh_results'');
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1984
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1985
                    (** cs # fK as ts rs , cs # fK bs us vs **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1986
                    val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1987
                    fun prove_fresh_result' recs t (a as Free (_, aT)) =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1988
                      Goal.prove context'' [] []
25823
5d75f4b179e2 Added function fresh_const.
berghofe
parents: 25674
diff changeset
  1989
                        (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1990
                        (fn _ => EVERY
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1991
                          [cut_facts_tac recs 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1992
                           REPEAT_DETERM (dresolve_tac context''
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1993
                             (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1994
                           NominalPermeq.fresh_guess_tac
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  1995
                             (put_simpset HOL_ss context'' addsimps (freshs2 @
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1996
                                fs_atoms @ fresh_atm @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  1997
                                maps snd finite_thss)) 1]);
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1998
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  1999
                    val fresh_results' =
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2000
                      map (prove_fresh_result' rec_prems1 rhs') freshs1 @
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2001
                      map (prove_fresh_result' rec_prems2 lhs') freshs1;
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2002
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2003
                    (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2004
                    val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2005
                    val pi1_pi2_result = Goal.prove context'' [] []
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2006
                      (HOLogic.mk_Trueprop (HOLogic.mk_eq
22311
ebcd44cb8d61 Curried and exported mk_perm.
berghofe
parents: 22274
diff changeset
  2007
                        (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  2008
                      (fn _ => simp_tac (put_simpset HOL_ss context''
27450
d45d2850aaed Replaced all but one occurrence of perm_simp_tac by perm_simproc_app,
berghofe
parents: 27300
diff changeset
  2009
                           addsimps pi1_pi2_eqs @ rec_eqns
d45d2850aaed Replaced all but one occurrence of perm_simp_tac by perm_simproc_app,
berghofe
parents: 27300
diff changeset
  2010
                           addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  2011
                         TRY (simp_tac (put_simpset HOL_ss context'' addsimps
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2012
                           (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2013
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2014
                    val _ = warning "final result";
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59499
diff changeset
  2015
                    val final = Goal.prove context'' [] [] (Thm.term_of concl)
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2016
                      (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
  2017
                        full_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh @
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2018
                          fresh_results @ fresh_results') 1);
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41489
diff changeset
  2019
                    val final' = Proof_Context.export context'' context' [final];
20376
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2020
                    val _ = warning "finished!"
53b31f7c1d87 Finished implementation of uniqueness proof for recursion combinator.
berghofe
parents: 20267
diff changeset
  2021
                  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  2022
                    resolve_tac context' final' 1
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2023
                  end) context 1])) idxss) (ndescr ~~ rec_elims))
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2024
         end));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2025
46161
4ed94d92ae19 prefer antiquotations;
wenzelm
parents: 45909
diff changeset
  2026
    val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2027
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2028
    (* define primrec combinators *)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2029
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45654
diff changeset
  2030
    val big_reccomb_name = space_implode "_" new_type_names ^ "_rec";
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28736
diff changeset
  2031
    val reccomb_names = map (Sign.full_bname thy11)
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2032
      (if length descr'' = 1 then [big_reccomb_name] else
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2033
        (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2034
          (1 upto (length descr''))));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2035
    val reccombs = map (fn ((name, T), T') => list_comb
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2036
      (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2037
        (reccomb_names ~~ recTs ~~ rec_result_Ts);
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2038
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2039
    val (reccomb_defs, thy12) =
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2040
      thy11
56239
17df7145a871 tuned signature;
wenzelm
parents: 55990
diff changeset
  2041
      |> Sign.add_consts (map (fn ((name, T), T') =>
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
  2042
          (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn))
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2043
          (reccomb_names ~~ recTs ~~ rec_result_Ts))
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  2044
      |> (Global_Theory.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
  2045
          (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T)
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
  2046
           (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
  2047
             (set $ Free ("x", T) $ Free ("y", T'))))))
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
  2048
               (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2049
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2050
    (* prove characteristic equations for primrec combinators *)
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2051
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2052
    val rec_thms = map (fn (prems, concl) =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2053
      let
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2054
        val _ $ (_ $ (_ $ x) $ _) = concl;
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2055
        val (_, cargs) = strip_comb x;
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2056
        val ps = map (fn (x as Free (_, T), i) =>
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2057
          (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2058
        val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  2059
        val prems' = flat finite_premss @ finite_ctxt_prems @
21088
13348ab97f5a Added freshness context to FCBs.
berghofe
parents: 21073
diff changeset
  2060
          rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  2061
        fun solve ctxt rules prems = resolve_tac ctxt rules THEN_ALL_NEW
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2062
          (resolve_tac ctxt prems THEN_ALL_NEW assume_tac ctxt)
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2063
      in
51122
386a117925db more parallel proofs in 'nominal_datatype', although sequential dark matter remains;
wenzelm
parents: 49835
diff changeset
  2064
        Goal.prove_global_future thy12 []
28731
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  2065
          (map (augment_sort thy12 fs_cp_sort) prems')
c60ac7923a06 Added support for parametric datatypes.
berghofe
parents: 28662
diff changeset
  2066
          (augment_sort thy12 fs_cp_sort concl')
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
  2067
          (fn {context = ctxt, prems} => EVERY
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
  2068
            [rewrite_goals_tac ctxt reccomb_defs,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2069
             resolve_tac ctxt @{thms the1_equality} 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  2070
             solve ctxt rec_unique_thms prems 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  2071
             resolve_tac ctxt rec_intrs 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  2072
             REPEAT (solve ctxt (prems @ rec_total_thms) prems 1)])
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2073
      end) (rec_eq_prems ~~
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  2074
        Old_Datatype_Prop.make_primrecs reccomb_names descr' thy12);
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
  2075
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45741
diff changeset
  2076
    val dt_infos = map_index (make_dt_info pdescr induct reccomb_names rec_thms)
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33040
diff changeset
  2077
      (descr1 ~~ distinct_thms ~~ inject_thms);
21540
f3faed8276e6 Additional information about nominal datatypes (descriptor,
berghofe
parents: 21516
diff changeset
  2078
19985
d39c554cf750 Implemented proofs of equivariance and finite support
berghofe
parents: 19874
diff changeset
  2079
    (* FIXME: theorems are stored in database for testing only *)
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2080
    val (_, thy13) = thy12 |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
  2081
      Global_Theory.add_thmss
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  2082
        [((Binding.name "rec_equiv", flat rec_equiv_thms), []),
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  2083
         ((Binding.name "rec_equiv'", flat rec_equiv_thms'), []),
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  2084
         ((Binding.name "rec_fin_supp", flat rec_fin_supp_thms), []),
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32727
diff changeset
  2085
         ((Binding.name "rec_fresh", flat rec_fresh_thms), []),
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33968
diff changeset
  2086
         ((Binding.name "rec_unique", map Drule.export_without_context rec_unique_thms), []),
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29389
diff changeset
  2087
         ((Binding.name "recs", rec_thms), [])] ||>
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24459
diff changeset
  2088
      Sign.parent_path ||>
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
  2089
      map_nominal_datatypes (fold Symtab.update dt_infos);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2090
  in
20397
243293620225 - Fixed bug that caused uniqueness proof for recursion
berghofe
parents: 20376
diff changeset
  2091
    thy13
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2092
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2093
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  2094
val nominal_datatype = gen_nominal_datatype Old_Datatype.check_specs;
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  2095
val nominal_datatype_cmd = gen_nominal_datatype Old_Datatype.read_specs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2096
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24814
diff changeset
  2097
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59880
diff changeset
  2098
  Outer_Syntax.command @{command_keyword nominal_datatype} "define nominal datatypes"
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  2099
    (Parse.and_list1 Old_Datatype.spec_cmd >>
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58111
diff changeset
  2100
      (Toplevel.theory o nominal_datatype_cmd Old_Datatype_Aux.default_config));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2101
18261
1318955d57ac Corrected treatment of non-recursive abstraction types.
berghofe
parents: 18246
diff changeset
  2102
end