Nominal/Perm.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 24 Jun 2010 19:32:33 +0100
changeset 2335 558c823f96aa
parent 2305 93ab397f5980
child 2337 b151399bd2c3
permissions -rw-r--r--
fixed according to changes in quotient
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1159
3c6bee89d826 Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory Perm
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
     2
imports 
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
     3
  "../Nominal-General/Nominal2_Base"
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
     4
  "../Nominal-General/Nominal2_Atoms" 
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
     5
  "../Nominal-General/Nominal2_Eqvt" 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
     6
  "Nominal2_FSet"
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
     7
  "Abs"
2305
93ab397f5980 smaller code for raw-eqvt proofs
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
     8
uses ("nominal_dt_rawperm.ML")
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
     9
     ("nominal_dt_rawfuns.ML")
2305
93ab397f5980 smaller code for raw-eqvt proofs
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    10
     ("nominal_dt_alpha.ML")
1159
3c6bee89d826 Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
begin
3c6bee89d826 Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    13
use "nominal_dt_rawperm.ML"
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    14
ML {* open Nominal_Dt_RawPerm *}
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    15
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    16
use "nominal_dt_rawfuns.ML"
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    17
ML {* open Nominal_Dt_RawFuns *}
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    18
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    19
use "nominal_dt_alpha.ML"
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    20
ML {* open Nominal_Dt_Alpha *}
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    21
1910
57891245370d tuning of the code
Christian Urban <urbanc@in.tum.de>
parents: 1903
diff changeset
    22
(* permutations for quotient types *)
1899
8e0bfb14f6bf optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents: 1896
diff changeset
    23
1253
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    24
ML {*
1903
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    25
fun quotient_lift_consts_export qtys spec ctxt =
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    26
let
2335
558c823f96aa fixed according to changes in quotient
Christian Urban <urbanc@in.tum.de>
parents: 2305
diff changeset
    27
  val (result, ctxt') = fold_map (Quotient_Def.lift_raw_const qtys) spec ctxt;
558c823f96aa fixed according to changes in quotient
Christian Urban <urbanc@in.tum.de>
parents: 2305
diff changeset
    28
  val (ts_loc, defs_loc) = split_list (map (fn info => (#qconst info, #def info)) result);
1903
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    29
  val morphism = ProofContext.export_morphism ctxt' ctxt;
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    30
  val ts = map (Morphism.term morphism) ts_loc
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    31
  val defs = Morphism.fact morphism defs_loc
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    32
in
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    33
  (ts, defs, ctxt')
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    34
end
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    35
*}
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    36
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    37
ML {*
1683
f78c820f67c3 Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1503
diff changeset
    38
fun define_lifted_perms qtys full_tnames name_term_pairs thms thy =
1253
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    39
let
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    40
  val lthy =
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    41
    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: 1503
diff changeset
    42
  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: 1503
diff changeset
    43
  val lifted_thms = map (Quotient_Tacs.lifted qtys lthy') thms;
1253
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    44
  fun tac _ =
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    45
    Class.intro_classes_tac [] THEN
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    46
    (ALLGOALS (resolve_tac lifted_thms))
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    47
  val lthy'' = Class.prove_instantiation_instance tac lthy'
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    48
in
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    49
  Local_Theory.exit_global lthy''
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    50
end
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    51
*}
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    52
1342
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    53
1899
8e0bfb14f6bf optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents: 1896
diff changeset
    54
end