| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 01 Apr 2010 17:56:26 +0200 | |
| changeset 1761 | 6bf14c13c291 | 
| parent 1683 | f78c820f67c3 | 
| child 1774 | c34347ec7ab3 | 
| permissions | -rw-r--r-- | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 1 | theory Perm | 
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 2 | imports "Nominal2_Atoms" | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | |
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 5 | ML {*
 | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 6 | open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *) | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 7 |   fun permute ty = Const (@{const_name permute}, @{typ perm} --> ty --> ty);
 | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 8 |   val minus_perm = Const (@{const_name minus}, @{typ perm} --> @{typ perm});
 | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 9 | *} | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 10 | |
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 11 | ML {*
 | 
| 1683 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 12 | fun quotient_lift_consts_export qtys spec ctxt = | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 13 | let | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 14 | val (result, ctxt') = fold_map (Quotient_Def.quotient_lift_const qtys) spec ctxt; | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 15 | val (ts_loc, defs_loc) = split_list result; | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 16 | val morphism = ProofContext.export_morphism ctxt' ctxt; | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 17 | val ts = map (Morphism.term morphism) ts_loc | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 18 | val defs = Morphism.fact morphism defs_loc | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 19 | in | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 20 | (ts, defs, ctxt') | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 21 | end | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 22 | *} | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 23 | |
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 24 | ML {*
 | 
| 1248 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 25 | fun prove_perm_empty lthy induct perm_def perm_frees = | 
| 1247 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 26 | let | 
| 1248 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 27 | val perm_types = map fastype_of perm_frees; | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 28 | val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types); | 
| 1256 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 29 | fun glc ((perm, T), x) = | 
| 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 30 |     HOLogic.mk_eq (perm $ @{term "0 :: perm"} $ Free (x, T), Free (x, T))
 | 
| 1247 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 31 | val gl = | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 32 | HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj | 
| 1256 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 33 | (map glc (perm_frees ~~ map body_type perm_types ~~ perm_indnames))); | 
| 1247 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 34 | fun tac _ = | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 35 | EVERY [ | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 36 | indtac induct perm_indnames 1, | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 37 |       ALLGOALS (asm_full_simp_tac (HOL_ss addsimps (@{thm permute_zero} :: perm_def)))
 | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 38 | ]; | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 39 | in | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 40 | split_conj_thm (Goal.prove lthy perm_indnames [] gl tac) | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 41 | end; | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 42 | *} | 
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 43 | |
| 1248 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 44 | ML {*
 | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 45 | fun prove_perm_append lthy induct perm_def perm_frees = | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 46 | let | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 47 |   val add_perm = @{term "op + :: (perm \<Rightarrow> perm \<Rightarrow> perm)"}
 | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 48 |   val pi1 = Free ("pi1", @{typ perm});
 | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 49 |   val pi2 = Free ("pi2", @{typ perm});
 | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 50 | val perm_types = map fastype_of perm_frees | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 51 | val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types); | 
| 1256 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 52 | fun glc ((perm, T), x) = | 
| 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 53 | HOLogic.mk_eq ( | 
| 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 54 | perm $ (add_perm $ pi1 $ pi2) $ Free (x, T), | 
| 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 55 | perm $ pi1 $ (perm $ pi2 $ Free (x, T))) | 
| 1248 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 56 | val gl = | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 57 | (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj | 
| 1256 
6c938f84880c
Restructuring the code in Perm
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1253diff
changeset | 58 | (map glc (perm_frees ~~ map body_type perm_types ~~ perm_indnames)))) | 
| 1248 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 59 | fun tac _ = | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 60 | EVERY [ | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 61 | indtac induct perm_indnames 1, | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 62 |       ALLGOALS (asm_full_simp_tac (HOL_ss addsimps (@{thm permute_plus} :: perm_def)))
 | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 63 | ] | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 64 | in | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 65 |   split_conj_thm (Goal.prove lthy ("pi1" :: "pi2" :: perm_indnames) [] gl tac)
 | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 66 | end; | 
| 
705afaaf6fb4
More refactoring and removed references to the global simpset in Perm.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1247diff
changeset | 67 | *} | 
| 1247 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 68 | |
| 
a728e199851d
Factor-out 'prove_perm_empty'; I plan to use it in defining permutations on the lifted type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1189diff
changeset | 69 | ML {*
 | 
| 1277 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1259diff
changeset | 70 | fun define_raw_perms (dt_info : info) number thy = | 
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 71 | let | 
| 1277 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1259diff
changeset | 72 |   val {descr, induct, sorts, ...} = dt_info;
 | 
| 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1259diff
changeset | 73 | val all_full_tnames = map (fn (_, (n, _, _)) => n) descr; | 
| 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1259diff
changeset | 74 | val full_tnames = List.take (all_full_tnames, number); | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 76 | val perm_names' = Datatype_Prop.indexify_names (map (fn (i, _) => | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 77 | "permute_" ^ name_of_typ (nth_dtyp i)) descr); | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 78 | val perm_types = map (fn (i, _) => | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | let val T = nth_dtyp i | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 |     in @{typ perm} --> T --> T end) descr;
 | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 | val perm_names_types' = perm_names' ~~ perm_types; | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 82 |   val pi = Free ("pi", @{typ perm});
 | 
| 1164 | 83 | fun perm_eq_constr i (cname, dts) = | 
| 84 | let | |
| 85 | val Ts = map (typ_of_dtyp descr sorts) dts; | |
| 86 | val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); | |
| 87 | val args = map Free (names ~~ Ts); | |
| 88 | val c = Const (cname, Ts ---> (nth_dtyp i)); | |
| 89 | fun perm_arg (dt, x) = | |
| 90 | let val T = type_of x | |
| 91 | in | |
| 92 | if is_rec_type dt then | |
| 93 | let val (Us, _) = strip_type T | |
| 94 | in list_abs (map (pair "x") Us, | |
| 95 | Free (nth perm_names_types' (body_index dt)) $ pi $ | |
| 96 | list_comb (x, map (fn (i, U) => | |
| 97 | (permute U) $ (minus_perm $ pi) $ Bound i) | |
| 98 | ((length Us - 1 downto 0) ~~ Us))) | |
| 99 | end | |
| 100 | else (permute T) $ pi $ x | |
| 101 | end; | |
| 102 | in | |
| 103 | (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq | |
| 104 | (Free (nth perm_names_types' i) $ | |
| 105 |            Free ("pi", @{typ perm}) $ list_comb (c, args),
 | |
| 106 | list_comb (c, map perm_arg (dts ~~ args))))) | |
| 107 | end; | |
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 108 | fun perm_eq (i, (_, _, constrs)) = map (perm_eq_constr i) constrs; | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 109 | val perm_eqs = maps perm_eq descr; | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 110 | val lthy = | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 111 |       Theory_Target.instantiation (full_tnames, [], @{sort pt}) thy;
 | 
| 1257 | 112 | val ((perm_frees, perm_ldef), lthy') = | 
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 113 | Primrec.add_primrec | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 114 | (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs lthy; | 
| 1277 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1259diff
changeset | 115 | val perm_empty_thms = List.take (prove_perm_empty lthy' induct perm_ldef perm_frees, number); | 
| 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1259diff
changeset | 116 | val perm_append_thms = List.take (prove_perm_append lthy' induct perm_ldef perm_frees, number) | 
| 1249 
ea6a52a4f5bf
Note the instance proofs, since they can be easily lifted.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1248diff
changeset | 117 | val perms_name = space_implode "_" perm_names' | 
| 
ea6a52a4f5bf
Note the instance proofs, since they can be easily lifted.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1248diff
changeset | 118 | val perms_zero_bind = Binding.name (perms_name ^ "_zero") | 
| 
ea6a52a4f5bf
Note the instance proofs, since they can be easily lifted.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1248diff
changeset | 119 | val perms_append_bind = Binding.name (perms_name ^ "_append") | 
| 1257 | 120 | fun tac _ (_, simps, _) = | 
| 1253 | 121 | (Class.intro_classes_tac []) THEN (ALLGOALS (resolve_tac simps)); | 
| 1257 | 122 | fun morphism phi (dfs, simps, fvs) = | 
| 123 | (map (Morphism.thm phi) dfs, map (Morphism.thm phi) simps, map (Morphism.term phi) fvs); | |
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 124 | in | 
| 1249 
ea6a52a4f5bf
Note the instance proofs, since they can be easily lifted.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1248diff
changeset | 125 | lthy' | 
| 
ea6a52a4f5bf
Note the instance proofs, since they can be easily lifted.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1248diff
changeset | 126 | |> snd o (Local_Theory.note ((perms_zero_bind, []), perm_empty_thms)) | 
| 
ea6a52a4f5bf
Note the instance proofs, since they can be easily lifted.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1248diff
changeset | 127 | |> snd o (Local_Theory.note ((perms_append_bind, []), perm_append_thms)) | 
| 1257 | 128 | |> Class_Target.prove_instantiation_exit_result morphism tac (perm_ldef, (perm_empty_thms @ perm_append_thms), perm_frees) | 
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 129 | end | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | |
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | *} | 
| 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 132 | |
| 1253 | 133 | ML {*
 | 
| 1683 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 134 | fun define_lifted_perms qtys full_tnames name_term_pairs thms thy = | 
| 1253 | 135 | let | 
| 136 | val lthy = | |
| 137 |     Theory_Target.instantiation (full_tnames, [], @{sort pt}) thy;
 | |
| 1683 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 138 | val (_, _, lthy') = quotient_lift_consts_export qtys name_term_pairs lthy; | 
| 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1503diff
changeset | 139 | val lifted_thms = map (Quotient_Tacs.lifted qtys lthy') thms; | 
| 1253 | 140 | fun tac _ = | 
| 141 | Class.intro_classes_tac [] THEN | |
| 142 | (ALLGOALS (resolve_tac lifted_thms)) | |
| 143 | val lthy'' = Class.prove_instantiation_instance tac lthy' | |
| 144 | in | |
| 145 | Local_Theory.exit_global lthy'' | |
| 146 | end | |
| 147 | *} | |
| 148 | ||
| 1342 | 149 | ML {*
 | 
| 150 | fun neq_to_rel r neq = | |
| 151 | let | |
| 152 | val neq = HOLogic.dest_Trueprop (prop_of neq) | |
| 153 | val eq = HOLogic.dest_not neq | |
| 154 | val (lhs, rhs) = HOLogic.dest_eq eq | |
| 155 | val rel = r $ lhs $ rhs | |
| 156 | val nrel = HOLogic.mk_not rel | |
| 157 | in | |
| 158 | HOLogic.mk_Trueprop nrel | |
| 159 | end | |
| 160 | *} | |
| 161 | ||
| 162 | ML {*
 | |
| 163 | fun neq_to_rel_tac cases distinct = | |
| 164 | rtac notI THEN' eresolve_tac cases THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps distinct) | |
| 165 | *} | |
| 166 | ||
| 167 | ML {*
 | |
| 168 | fun distinct_rel ctxt cases (dists, rel) = | |
| 169 | let | |
| 170 | val ((_, thms), ctxt') = Variable.import false dists ctxt | |
| 171 | val terms = map (neq_to_rel rel) thms | |
| 172 | val nrels = map (fn t => Goal.prove ctxt' [] [] t (fn _ => neq_to_rel_tac cases dists 1)) terms | |
| 173 | in | |
| 174 | Variable.export ctxt' ctxt nrels | |
| 175 | end | |
| 176 | *} | |
| 177 | ||
| 178 | ||
| 179 | ||
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 180 | (* Test | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 181 | atom_decl name | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 182 | |
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 183 | datatype rtrm1 = | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 184 | rVr1 "name" | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 185 | | rAp1 "rtrm1" "rtrm1 list" | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 186 | | rLm1 "name" "rtrm1" | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 187 | | rLt1 "bp" "rtrm1" "rtrm1" | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 188 | and bp = | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 189 | BUnit | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 190 | | BVr "name" | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 191 | | BPr "bp" "bp" | 
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 192 | |
| 1170 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 193 | |
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 194 | setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Perm.rtrm1", "Perm.bp"] *}
 | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 195 | print_theorems | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 196 | *) | 
| 
a7b4160ef463
Wrapped the permutation code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1164diff
changeset | 197 | |
| 1159 
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 198 | end |