Quot/Nominal/Fv.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 17 Feb 2010 15:00:04 +0100
changeset 1175 6a3be6ef348d
parent 1174 f6e9ae54b855
child 1176 29c4a0cf9237
permissions -rw-r--r--
Reorganization
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory Fv
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
1169
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
     5
(* Bindings are given as a list which has a length being equal
1172
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
     6
   to the length of the number of constructors.
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
     7
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
     8
   Each element is a list whose length is equal to the number
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
     9
   of arguents.
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    10
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    11
   Every element specifies bindings of this argument given as
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    12
   a tuple: function, bound argument.
1169
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    13
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    14
  Eg:
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    15
nominal_datatype
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    16
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    17
   C1
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    18
 | C2 x y z bind x in z
1172
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    19
 | C3 x y z bind f x in z bind g y in z
1169
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    20
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    21
yields:
1172
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    22
[
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    23
 [],
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    24
 [[], [], [(NONE, 0)]],
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    25
 [[], [], [(SOME (Const f), 0), (Some (Const g), 1)]]]
1169
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    26
*)
b9d02e0800e9 Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1168
diff changeset
    27
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
ML {*
1175
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    29
  open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *);
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    30
  fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom});
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    31
  val noatoms = @{term "{} :: atom set"};
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    32
  fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x];
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    33
  fun mk_union sets =
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    34
    fold (fn a => fn b =>
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    35
      if a = noatoms then b else
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    36
      if b = noatoms then a else
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    37
      HOLogic.mk_binop @{const_name union} (a, b)) (rev sets) noatoms;
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    38
  fun mk_diff a b =
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    39
    if b = noatoms then a else
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    40
    if b = a then noatoms else
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    41
    HOLogic.mk_binop @{const_name minus} (a, b);
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    42
*}
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    43
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    44
atom_decl name
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    45
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    46
datatype rlam =
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    47
  rVar "name"
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    48
| rApp "rlam" "rlam"
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    49
| rLam "name" "rlam"
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    50
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    51
ML {*
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
  val {descr, ...} = Datatype.the_info @{theory} "Fv.rlam";
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
  val sorts = [];
1175
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    54
  val binds = [[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]]];
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
  fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
  val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
    "fv_" ^ name_of_typ (nth_dtyp i)) descr);
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
  val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr;
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  val fv_frees = map Free (fv_names ~~ fv_types);
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    60
  fun fv_eq_constr i (cname, dts) bindcs =
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
    let
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
      val Ts = map (typ_of_dtyp descr sorts) dts;
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
      val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
      val args = map Free (names ~~ Ts);
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
      val c = Const (cname, Ts ---> (nth_dtyp i));
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
      val fv_c = Free (nth fv_names i, (nth_dtyp i) --> @{typ "atom set"});
1174
f6e9ae54b855 Now should work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1173
diff changeset
    67
      (* TODO we assume that all can be 'atomized' *)
f6e9ae54b855 Now should work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1173
diff changeset
    68
      fun fv_bind (NONE, i) = mk_single_atom (nth args i)
f6e9ae54b855 Now should work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1173
diff changeset
    69
        | fv_bind (SOME f, i) = f $ (nth args i);
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    70
      fun fv_arg ((dt, x), bindxs) =
1172
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    71
        let
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    72
          val arg =
1175
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    73
            if is_rec_type dt then nth fv_frees (body_index dt) $ x else
1174
f6e9ae54b855 Now should work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1173
diff changeset
    74
            (* TODO: we just assume everything can be 'atomized' *)
1175
6a3be6ef348d Reorganization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1174
diff changeset
    75
            HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    76
          val sub = mk_union (map fv_bind bindxs)
1172
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    77
        in
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    78
          mk_diff arg sub
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    79
        end;
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    80
    in
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
      (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    82
        (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ bindcs)))))
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
    end;
1172
9a609fefcf24 Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1169
diff changeset
    84
  fun fv_eq (i, (_, _, constrs)) = map2 (fv_eq_constr i) constrs binds;
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
  val fv_eqs = maps fv_eq descr
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
*}
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
1173
9cb99a28b40e Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1172
diff changeset
    88
1168
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
local_setup {*
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
snd o (Primrec.add_primrec
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
  (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) fv_eqs)
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
*}
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
print_theorems
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
fun
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
  fv_rlam :: "rlam \<Rightarrow> atom set"
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
where
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
  "fv_rlam (rVar a) = {atom a}"
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
| "fv_rlam (rApp t1 t2) = (fv_rlam t1) \<union> (fv_rlam t2)"
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
| "fv_rlam (rLam a t) = (fv_rlam t) - {atom a}"
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
5c1e16806901 Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
end