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