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