Nominal/NewFv.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 29 Apr 2010 16:15:49 +0200
changeset 1986 522748f37444
parent 1985 727e0edad284
child 1989 45721f92e471
permissions -rw-r--r--
Extracting the fv body function and exporting the terms.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory NewFv
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
     2
imports "../Nominal-General/Nominal2_Atoms"
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
     3
        "Abs" "Perm" "Nominal2_FSet"
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
begin
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
ML {*
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
datatype bmodes =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
   BEmy of int
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
|  BLst of ((term option * int) list) * (int list)
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
|  BSet of ((term option * int) list) * (int list)
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
|  BRes of ((term option * int) list) * (int list)
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
ML {*
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    15
fun mk_singleton_atom x = HOLogic.mk_set @{typ atom} [mk_atom x];
1962
84a13d1e2511 moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1960
diff changeset
    16
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    17
val noatoms = @{term "{} :: atom set"};
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    18
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    19
fun mk_union sets =
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    20
  fold (fn a => fn b =>
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    21
  if a = noatoms then b else
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    22
  if b = noatoms then a else
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    23
  if a = b then a else
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
    24
  HOLogic.mk_binop @{const_name sup} (a, b)) (rev sets) noatoms;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
ML {*
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    28
fun is_atom thy ty =
1970
90758c187861 use sort at_base instead of at
Christian Urban <urbanc@in.tum.de>
parents: 1969
diff changeset
    29
  Sign.of_sort thy (ty, @{sort at_base})
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
fun is_atom_set thy (Type ("fun", [t, @{typ bool}])) = is_atom thy t
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
  | is_atom_set _ _ = false;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
fun is_atom_fset thy (Type ("FSet.fset", [t])) = is_atom thy t
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  | is_atom_fset _ _ = false;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
fun mk_atom_set t =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  val ty = fastype_of t;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
  val atom_ty = HOLogic.dest_setT ty --> @{typ atom};
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
  val img_ty = atom_ty --> ty --> @{typ "atom set"};
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
in
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    43
  (Const (@{const_name image}, img_ty) $ mk_atom_ty atom_ty t)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
end;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
fun mk_atom_fset t =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
  val ty = fastype_of t;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
  val atom_ty = dest_fsetT ty --> @{typ atom};
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
  val fmap_ty = atom_ty --> ty --> @{typ "atom fset"};
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
  val fset_to_set = @{term "fset_to_set :: atom fset \<Rightarrow> atom set"}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
in
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
    53
  fset_to_set $ (Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
end;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
fun mk_diff a b =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
  if b = noatoms then a else
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
  if b = a then noatoms else
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  HOLogic.mk_binop @{const_name minus} (a, b);
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
ML {*
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    63
fun atoms thy t =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
let
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    65
  val ty = fastype_of t;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
in
1983
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    67
  if is_atom thy ty
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    68
    then mk_singleton_atom t
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    69
  else if is_atom_set thy ty
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    70
    then mk_atom_set t
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    71
  else if is_atom_fset thy ty
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    72
    then mk_atom_fset t
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    73
  else noatoms
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
ML {*
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
fun setify x =
1983
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    79
  if fastype_of x = @{typ "atom list"}
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    80
  then @{term "set::atom list \<Rightarrow> atom set"} $ x
1969
9ae5380c828a white spaces
Christian Urban <urbanc@in.tum.de>
parents: 1968
diff changeset
    81
  else x
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
ML {*
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    85
fun fv_body thy dts args fv_frees supp i =
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    86
let
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    87
  val x = nth args i;
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    88
  val dt = nth dts i;
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    89
in
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    90
  if Datatype_Aux.is_rec_type dt
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    91
  then nth fv_frees (Datatype_Aux.body_index dt) $ x
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    92
  else (if supp then mk_supp x else atoms thy x)
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    93
end
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    94
*}
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    95
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    96
ML {*
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
fun fv_bm_lsts thy dts args fv_frees bn_fvbn binds bodys =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
let
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
    99
  val fv_bodys = mk_union (map (fv_body thy dts args fv_frees true) bodys)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
  val bound_vars =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
    case binds of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
      [(SOME bn, i)] => setify (bn $ nth args i)
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   103
    | _ => mk_union (map (fv_body thy dts args fv_frees false) (map snd binds));
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
  val non_rec_vars =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
    case binds of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
      [(SOME bn, i)] =>
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
        if i mem bodys
1985
727e0edad284 Fix for recursive binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1984
diff changeset
   108
        then noatoms
727e0edad284 Fix for recursive binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1984
diff changeset
   109
        else ((the (AList.lookup (op=) bn_fvbn bn)) $ nth args i)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
    | _ => noatoms
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
  mk_union [mk_diff fv_bodys bound_vars, non_rec_vars]
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
ML {*
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
fun fv_bn_bm thy dts args fv_frees bn_fvbn args_in_bn bm =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
case bm of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
  BEmy i =>
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
    let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
      val x = nth args i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
      val dt = nth dts i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
    in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
      case AList.lookup (op=) args_in_bn i of
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   125
        NONE => if Datatype_Aux.is_rec_type dt
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   126
                then nth fv_frees (Datatype_Aux.body_index dt) $ x
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   127
                else mk_supp x
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
      | SOME (SOME (f : term)) => (the (AList.lookup (op=) bn_fvbn f)) $ x
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
      | SOME NONE => noatoms
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
    end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
| BLst (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
| BSet (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
| BRes (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
ML {*
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   137
fun fv_bn thy dt_descr sorts fv_frees bn_fvbn bclausess (fvbn, (_, ith_dtyp, args_in_bns)) =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
let
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   139
  fun fv_bn_constr (cname, dts) (args_in_bn, bclauses) =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
  let
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   141
    val Ts = map (Datatype_Aux.typ_of_dtyp dt_descr sorts) dts;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
    val names = Datatype_Prop.make_tnames Ts;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
    val args = map Free (names ~~ Ts);
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   144
    val c = Const (cname, Ts ---> (nth_dtyp dt_descr sorts ith_dtyp));
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
    val fv_bn_bm = fv_bn_bm thy dts args fv_frees bn_fvbn args_in_bn
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
  in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
    HOLogic.mk_Trueprop (HOLogic.mk_eq
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   148
      (fvbn $ list_comb (c, args), mk_union (map fv_bn_bm bclauses)))
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
  end;
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   150
  val (_, (_, _, constrs)) = nth dt_descr ith_dtyp;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
in
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   152
  map2 fv_bn_constr constrs (args_in_bns ~~ bclausess)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
ML {*
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   157
fun fv_bns thy dt_descr sorts fv_frees bn_funs bclausesss =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
  fun mk_fvbn_free (bn, ith, _) =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
    let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
      val fvbn_name = "fv_" ^ (Long_Name.base_name (fst (dest_Const bn)));
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
    in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
      (fvbn_name, Free (fvbn_name, fastype_of (nth fv_frees ith)))
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
    end;
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   165
  val (fvbn_names, fvbn_frees) = split_list (map mk_fvbn_free bn_funs);
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   166
  val bn_fvbn = (map (fn (bn, _, _) => bn) bn_funs) ~~ fvbn_frees
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   167
  val bclausessl = map (fn (_, i, _) => nth bclausesss i) bn_funs;
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   168
  val eqs = map2 (fv_bn thy dt_descr sorts fv_frees bn_fvbn) bclausessl (fvbn_frees ~~ bn_funs);
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
in
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   170
  (bn_fvbn, fvbn_names, eqs)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
ML {*
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
fun fv_bm thy dts args fv_frees bn_fvbn bm =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
case bm of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
  BEmy i =>
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
    let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
      val x = nth args i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
      val dt = nth dts i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
    in
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   182
      if Datatype_Aux.is_rec_type dt
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   183
      then nth fv_frees (Datatype_Aux.body_index dt) $ x
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   184
      else mk_supp x
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
    end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
| BLst (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
| BSet (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
| BRes (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
ML {*
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   192
fun fv thy dt_descr sorts fv_frees bn_fvbn bclausess (fv_free, ith_dtyp) =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
let
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   194
  fun fv_constr (cname, dts) bclauses =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
  let
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   196
    val Ts = map (Datatype_Aux.typ_of_dtyp dt_descr sorts) dts;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   197
    val names = Datatype_Prop.make_tnames Ts;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
    val args = map Free (names ~~ Ts);
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   199
    val c = Const (cname, Ts ---> (nth_dtyp dt_descr sorts ith_dtyp));
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
    val fv_bn_bm = fv_bm thy dts args fv_frees bn_fvbn
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   201
  in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   202
    HOLogic.mk_Trueprop (HOLogic.mk_eq
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   203
      (fv_free $ list_comb (c, args), mk_union (map fv_bn_bm bclauses)))
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
  end;
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   205
  val (_, (_, _, constrs)) = nth dt_descr ith_dtyp;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   206
in
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   207
  map2 fv_constr constrs bclausess
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   208
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   209
*}
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
   210
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   211
ML {*
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
   212
fun define_raw_fv (dt_info : Datatype_Aux.info) bn_funs bclausesss lthy =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   214
  val thy = ProofContext.theory_of lthy;
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   215
  val {descr as dt_descr, sorts, ...} = dt_info;
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   216
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   217
  val fv_names = prefix_dt_names dt_descr sorts "fv_"
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   218
  val fv_types = map (fn (i, _) => nth_dtyp descr sorts i --> @{typ "atom set"}) dt_descr;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   219
  val fv_frees = map Free (fv_names ~~ fv_types);
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   220
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   221
  val (bn_fvbn, fv_bn_names, fv_bn_eqs) =
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
   222
    fv_bns thy dt_descr sorts fv_frees bn_funs bclausesss;
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   223
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   224
  val fv_bns = map snd bn_fvbn;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   225
  val fv_nums = 0 upto (length fv_frees - 1)
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   226
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
   227
  val fv_eqs = map2 (fv thy dt_descr sorts fv_frees bn_fvbn) bclausesss (fv_frees ~~ fv_nums);
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   228
1971
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   229
  val all_fv_names = map (fn s => (Binding.name s, NONE, NoSyn)) (fv_names @ fv_bn_names)
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   230
  val all_fv_eqs = map (pair Attrib.empty_binding) (flat fv_eqs @ flat fv_bn_eqs)
1965
4a3c05fe2bc5 closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   231
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   232
  fun pat_completeness_auto ctxt =
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   233
    Pat_Completeness.pat_completeness_tac ctxt 1
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   234
      THEN auto_tac (clasimpset_of ctxt)
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   235
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   236
  fun prove_termination lthy =
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   237
    Function.prove_termination NONE
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   238
      (Lexicographic_Order.lexicographic_order_tac true lthy) lthy
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   239
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   240
  val (info, lthy') = Function.add_function all_fv_names all_fv_eqs
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   241
    Function_Common.default_config pat_completeness_auto lthy
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   242
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   243
  val lthy'' = prove_termination (Local_Theory.restore lthy')
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   244
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   245
  val morphism = ProofContext.export_morphism lthy'' lthy
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   246
  val fv_frees_exp = map (Morphism.term morphism) fv_frees
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   247
  val fv_bns_exp = map (Morphism.term morphism) fv_bns
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   248
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   249
in
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   250
  ((fv_frees_exp, fv_bns_exp), info, lthy'')
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   251
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   252
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   253
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   254
end