Nominal/NewFv.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 14 May 2010 10:21:14 +0200
changeset 2131 f7ec6f7b152e
parent 2071 843e0a2d44d7
child 2133 16834a4ca1bb
permissions -rw-r--r--
Fv for multiple binding functions
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 {*
2046
73c50e913db6 tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2000
diff changeset
     7
(* binding modes *)
73c50e913db6 tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2000
diff changeset
     8
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
datatype bmodes =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
   BEmy of int
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
|  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
    12
|  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
    13
|  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
    14
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
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
    17
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
    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
val noatoms = @{term "{} :: atom set"};
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    20
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
    21
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
    22
  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
    23
  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
    24
  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
    25
  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
    26
  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
    27
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
ML {*
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    30
fun is_atom thy ty =
1970
90758c187861 use sort at_base instead of at
Christian Urban <urbanc@in.tum.de>
parents: 1969
diff changeset
    31
  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
    32
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
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
    34
  | is_atom_set _ _ = false;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
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
    37
  | is_atom_fset _ _ = false;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
fun mk_atom_set t =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
  val ty = fastype_of t;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
  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
    43
  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
    44
in
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    45
  (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
    46
end;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
fun mk_atom_fset t =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
  val ty = fastype_of t;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
  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
    52
  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
    53
  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
    54
in
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
    55
  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
    56
end;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
fun mk_diff a b =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  if b = noatoms then a else
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
  if b = a then noatoms else
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
  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
    62
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
ML {*
1989
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    65
fun is_atom_list (Type (@{type_name list}, [T])) = true
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    66
  | is_atom_list _ = false
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    67
*}
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    68
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    69
ML {*
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    70
fun dest_listT (Type (@{type_name list}, [T])) = T
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    71
  | dest_listT T = raise TYPE ("dest_listT: list type expected", [T], [])
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    72
*}
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    73
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    74
ML {*
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    75
fun mk_atom_list t =
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    76
let
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    77
  val ty = fastype_of t;
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    78
  val atom_ty = dest_listT ty --> @{typ atom};
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    79
  val map_ty = atom_ty --> ty --> @{typ "atom list"};
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    80
in
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    81
  (Const (@{const_name map}, map_ty) $ mk_atom_ty atom_ty t)
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    82
end;
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    83
*}
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    84
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    85
ML {*
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    86
fun setify thy t =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
let
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
    88
  val ty = fastype_of t;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
in
1983
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
    90
  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
    91
    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
    92
  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
    93
    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
    94
  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
    95
    then mk_atom_fset t
1989
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
    96
  else error ("setify" ^ (PolyML.makestring (t, ty)))
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
ML {*
1989
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   101
fun listify thy t =
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   102
let
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   103
  val ty = fastype_of t;
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   104
in
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   105
  if is_atom thy ty
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   106
    then HOLogic.mk_list @{typ atom} [mk_atom t]
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   107
  else if is_atom_list ty
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   108
    then mk_atom_set t
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   109
  else error "listify"
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   110
end
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   111
*}
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   112
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   113
ML {*
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   114
fun set x =
1983
4538d63ecc9b Support in positive position and atoms in negative positions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   115
  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
   116
  then @{term "set::atom list \<Rightarrow> atom set"} $ x
1969
9ae5380c828a white spaces
Christian Urban <urbanc@in.tum.de>
parents: 1968
diff changeset
   117
  else x
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
ML {*
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   121
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
   122
let
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   123
  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
   124
  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
   125
in
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   126
  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
   127
  then nth fv_frees (Datatype_Aux.body_index dt) $ x
1989
45721f92e471 Unify and give only one name to 'setify', 'listify' and 'set'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1986
diff changeset
   128
  else (if supp then mk_supp x else setify thy x)
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   129
end
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   130
*}
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   131
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   132
ML {*
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
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
   134
let
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   135
  val fv_bodys = mk_union (map (fv_body thy dts args fv_frees true) bodys)
2131
f7ec6f7b152e Fv for multiple binding functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2071
diff changeset
   136
  fun bound_var (SOME bn, i) = set (bn $ nth args i)
f7ec6f7b152e Fv for multiple binding functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2071
diff changeset
   137
    | bound_var (NONE, i) = fv_body thy dts args fv_frees false i
f7ec6f7b152e Fv for multiple binding functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2071
diff changeset
   138
  val bound_vars = mk_union (map bound_var binds);
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
  val non_rec_vars =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
    case binds of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
      [(SOME bn, i)] =>
2071
843e0a2d44d7 mem => member
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2046
diff changeset
   142
        if member (op =) bodys i
1985
727e0edad284 Fix for recursive binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1984
diff changeset
   143
        then noatoms
727e0edad284 Fix for recursive binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1984
diff changeset
   144
        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
   145
    | _ => noatoms
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
  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
   148
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
ML {*
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
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
   153
case bm of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
  BEmy i =>
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
    let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
      val x = nth args i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
      val dt = nth dts i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
    in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
      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
   160
        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
   161
                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
   162
                else mk_supp x
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
      | 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
   164
      | SOME NONE => noatoms
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
    end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
| 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
   167
| 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
   168
| 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
   169
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
ML {*
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   172
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
   173
let
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   174
  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
   175
  let
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   176
    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
   177
    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
   178
    val args = map Free (names ~~ Ts);
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   179
    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
   180
    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
   181
  in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
    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
   183
      (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
   184
  end;
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   185
  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
   186
in
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   187
  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
   188
end
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_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
   193
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
  fun mk_fvbn_free (bn, ith, _) =
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   195
    let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   196
      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
   197
    in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   198
      (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
   199
    end;
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   200
  val (fvbn_names, fvbn_frees) = split_list (map mk_fvbn_free bn_funs);
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   201
  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
   202
  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
   203
  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
   204
in
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   205
  (bn_fvbn, fvbn_names, eqs)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   206
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   207
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   208
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   209
ML {*
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   210
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
   211
case bm of
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   212
  BEmy i =>
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 x = nth args i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   215
      val dt = nth dts i;
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   216
    in
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   217
      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
   218
      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
   219
      else mk_supp x
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   220
    end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   221
| 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
   222
| 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
   223
| 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
   224
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   225
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   226
ML {*
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   227
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
   228
let
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   229
  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
   230
  let
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   231
    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
   232
    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
   233
    val args = map Free (names ~~ Ts);
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   234
    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
   235
    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
   236
  in
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   237
    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
   238
      (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
   239
  end;
1968
98303b78fb64 avoided repeated dest of dt_info
Christian Urban <urbanc@in.tum.de>
parents: 1967
diff changeset
   240
  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
   241
in
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   242
  map2 fv_constr constrs bclausess
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   243
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   244
*}
1963
0c9ef14e9ba4 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1962
diff changeset
   245
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   246
ML {*
2046
73c50e913db6 tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2000
diff changeset
   247
fun define_raw_fv dt_descr sorts bn_funs bclausesss lthy =
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   248
let
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   249
  val thy = ProofContext.theory_of lthy;
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
   250
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   251
  val fv_names = prefix_dt_names dt_descr sorts "fv_"
2046
73c50e913db6 tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2000
diff changeset
   252
  val fv_types = map (fn (i, _) => nth_dtyp dt_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
   253
  val fv_frees = map Free (fv_names ~~ fv_types);
1967
Christian Urban <urbanc@in.tum.de>
parents: 1966
diff changeset
   254
2046
73c50e913db6 tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2000
diff changeset
   255
  (* free variables for the bn-functions *)
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   256
  val (bn_fvbn, fv_bn_names, fv_bn_eqs) =
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
   257
    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
   258
2046
73c50e913db6 tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2000
diff changeset
   259
  
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   260
  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
   261
  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
   262
1984
92f40c13d581 revert 0c9ef14e9ba4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1983
diff changeset
   263
  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
   264
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
   265
  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
   266
  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
   267
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   268
  fun pat_completeness_auto ctxt =
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   269
    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
   270
      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
   271
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   272
  fun prove_termination lthy =
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   273
    Function.prove_termination NONE
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   274
      (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
   275
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   276
  val (_, lthy') = Function.add_function all_fv_names all_fv_eqs
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   277
    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
   278
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   279
  val (info, lthy'') = prove_termination (Local_Theory.restore lthy')
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   280
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   281
  val {fs, simps, ...} = info;
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   282
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   283
  val morphism = ProofContext.export_morphism lthy'' lthy
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   284
  val fs_exp = map (Morphism.term morphism) fs
1986
522748f37444 Extracting the fv body function and exporting the terms.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1985
diff changeset
   285
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   286
  val (fv_frees_exp, fv_bns_exp) = chop (length fv_frees) fs_exp
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   287
  val simps_exp = Morphism.fact morphism (the simps)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   288
in
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   289
  ((fv_frees_exp, fv_bns_exp), simps_exp, lthy'')
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   290
end
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   291
*}
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   292
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   293
end