Nominal/Perm.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 24 May 2010 21:11:33 +0100
changeset 2298 aead2aad845c
parent 2297 9ca7b249760e
child 2300 9fb315392493
permissions -rw-r--r--
tuned
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" 
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
     5
  "Nominal2_FSet" "Abs"
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
     6
uses ("nominal_dt_rawperm.ML")
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
     7
     ("nominal_dt_rawfuns.ML")
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
     8
     ("nominal_dt_alpha.ML")
1159
3c6bee89d826 Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
begin
3c6bee89d826 Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    11
use "nominal_dt_rawperm.ML"
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    12
ML {* open Nominal_Dt_RawPerm *}
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    13
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    14
use "nominal_dt_rawfuns.ML"
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    15
ML {* open Nominal_Dt_RawFuns *}
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    16
2297
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    17
use "nominal_dt_alpha.ML"
Christian Urban <urbanc@in.tum.de>
parents: 2296
diff changeset
    18
ML {* open Nominal_Dt_Alpha *}
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2288
diff changeset
    19
1910
57891245370d tuning of the code
Christian Urban <urbanc@in.tum.de>
parents: 1903
diff changeset
    20
(* permutations for quotient types *)
1899
8e0bfb14f6bf optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents: 1896
diff changeset
    21
1253
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    22
ML {*
1903
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    23
fun quotient_lift_consts_export qtys spec ctxt =
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    24
let
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    25
  val (result, ctxt') = fold_map (Quotient_Def.quotient_lift_const qtys) spec ctxt;
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    26
  val (ts_loc, defs_loc) = split_list result;
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    27
  val morphism = ProofContext.export_morphism ctxt' ctxt;
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    28
  val ts = map (Morphism.term morphism) ts_loc
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    29
  val defs = Morphism.fact morphism defs_loc
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    30
in
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    31
  (ts, defs, ctxt')
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    32
end
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    33
*}
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    34
950fd9b8f05e reordered code
Christian Urban <urbanc@in.tum.de>
parents: 1902
diff changeset
    35
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
    36
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
    37
let
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    38
  val lthy =
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    39
    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
    40
  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
    41
  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
    42
  fun tac _ =
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    43
    Class.intro_classes_tac [] THEN
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    44
    (ALLGOALS (resolve_tac lifted_thms))
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    45
  val lthy'' = Class.prove_instantiation_instance tac lthy'
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    46
in
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    47
  Local_Theory.exit_global lthy''
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    48
end
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    49
*}
cff8a67691d2 Define lifted perms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1249
diff changeset
    50
1342
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    51
ML {*
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    52
fun neq_to_rel r neq =
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    53
let
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    54
  val neq = HOLogic.dest_Trueprop (prop_of neq)
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    55
  val eq = HOLogic.dest_not neq
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    56
  val (lhs, rhs) = HOLogic.dest_eq eq
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    57
  val rel = r $ lhs $ rhs
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    58
  val nrel = HOLogic.mk_not rel
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    59
in
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    60
  HOLogic.mk_Trueprop nrel
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    61
end
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    62
*}
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    63
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    64
ML {*
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    65
fun neq_to_rel_tac cases distinct =
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    66
  rtac notI THEN' eresolve_tac cases THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps distinct)
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    67
*}
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    68
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    69
ML {*
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    70
fun distinct_rel ctxt cases (dists, rel) =
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    71
let
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    72
  val ((_, thms), ctxt') = Variable.import false dists ctxt
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    73
  val terms = map (neq_to_rel rel) thms
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    74
  val nrels = map (fn t => Goal.prove ctxt' [] [] t (fn _ => neq_to_rel_tac cases dists 1)) terms
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    75
in
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    76
  Variable.export ctxt' ctxt nrels
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    77
end
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    78
*}
2b98012307f7 Lift distinct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    79
1899
8e0bfb14f6bf optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents: 1896
diff changeset
    80
end