Nominal/NewParser.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 20 May 2010 21:23:53 +0100
changeset 2288 3b83960f9544
parent 2146 a2f70c09e77d
child 2292 d134bd4f6d1b
permissions -rw-r--r--
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory NewParser
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
     2
imports "../Nominal-General/Nominal2_Base" 
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
        "../Nominal-General/Nominal2_Eqvt" 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
        "../Nominal-General/Nominal2_Supp" 
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
     5
        "Perm" "NewAlpha" "Tacs" "Equivp" "Lift"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
begin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
     8
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
section{* Interface for nominal_datatype *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
(* nominal datatype parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
local
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  structure P = OuterParse;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  structure S = Scan
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  fun triple1 ((x, y), z) = (x, y, z)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  fun triple2 (x, (y, z)) = (x, y, z)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  fun tuple ((x, y, z), u) = (x, y, z, u)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  fun tswap (((x, y), z), u) = (x, y, u, z)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
val _ = OuterKeyword.keyword "bind"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
val _ = OuterKeyword.keyword "bind_set"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
val _ = OuterKeyword.keyword "bind_res"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
    28
val anno_typ = S.option (P.name --| P.$$$ "::") -- P.typ
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
val bind_mode = P.$$$ "bind" || P.$$$ "bind_set" || P.$$$ "bind_res"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
val bind_clauses = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  P.enum "," (bind_mode -- S.repeat1 P.term -- (P.$$$ "in" |-- S.repeat1 P.name) >> triple1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
val cnstr_parser =
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
    36
  P.binding -- S.repeat anno_typ -- bind_clauses -- P.opt_mixfix >> tswap
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
(* datatype parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
val dt_parser =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  (P.type_args -- P.binding -- P.opt_mixfix >> triple1) -- 
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
    41
    (P.$$$ "=" |-- P.enum1 "|" cnstr_parser) >> tuple
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
(* binding function parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
val bnfun_parser = 
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    45
  S.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([], [])
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
(* main parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
val main_parser =
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    49
  P.and_list1 dt_parser -- bnfun_parser >> triple2
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
fun get_cnstrs dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  map (fn (_, _, _, constrs) => constrs) dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
fun get_typed_cnstrs dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  flat (map (fn (_, bn, _, constrs) => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
   (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
fun get_cnstr_strs dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  map (fn (bn, _, _) => Binding.name_of bn) (flat (get_cnstrs dts))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
fun get_bn_fun_strs bn_funs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
    69
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
ML {* 
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
    71
(* exports back the results *)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
fun add_primrec_wrapper funs eqs lthy = 
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
    73
  if null funs then ([], [], lthy)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  else 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
   let 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
     val eqs' = map (fn (_, eq) => (Attrib.empty_binding, eq)) eqs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
     val funs' = map (fn (bn, ty, mx) => (bn, SOME ty, mx)) funs
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
    78
     val ((funs'', eqs''), lthy') = Primrec.add_primrec funs' eqs' lthy
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
    79
     val phi = ProofContext.export_morphism lthy' lthy
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
   in 
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
    81
     (map (Morphism.term phi) funs'', map (Morphism.thm phi) eqs'', lthy')
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
   end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
fun add_datatype_wrapper dt_names dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  val conf = Datatype.default_config
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
  Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    94
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    95
text {* Infrastructure for adding "_raw" to types and terms *}
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    96
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
fun add_raw s = s ^ "_raw"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
fun add_raws ss = map add_raw ss
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
fun raw_bind bn = Binding.suffix_name "_raw" bn
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
fun replace_str ss s = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  case (AList.lookup (op=) ss s) of 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
     SOME s' => s'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
   | NONE => s
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  | replace_typ ty_ss T = T  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
fun raw_dts ty_ss dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  fun raw_dts_aux1 (bind, tys, mx) =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
    (raw_bind bind, map (replace_typ ty_ss) tys, mx)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
    (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
  map raw_dts_aux2 dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  | replace_aterm trm_ss trm = trm
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
fun replace_term trm_ss ty_ss trm =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
fun rawify_dts dt_names dts dts_env =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  val raw_dts = raw_dts dts_env dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  val raw_dt_names = add_raws dt_names
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  (raw_dt_names, raw_dts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
fun rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
  val bn_funs' = map (fn (bn, ty, mx) => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
    (raw_bind bn, replace_typ dts_env ty, mx)) bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  val bn_eqs' = map (fn (attr, trm) => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
    (attr, replace_term (cnstrs_env @ bn_fun_env) dts_env trm)) bn_eqs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
  (bn_funs', bn_eqs') 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
fun rawify_bclauses dts_env cnstrs_env bn_fun_env bclauses =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  fun rawify_bnds bnds = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
    map (apfst (Option.map (replace_term (cnstrs_env @ bn_fun_env) dts_env))) bnds
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   158
  fun rawify_bclause (BC (mode, bnds, bdys)) = BC (mode, rawify_bnds bnds, bdys)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  map (map (map rawify_bclause)) bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   164
(* strip_bn_fun takes a rhs of a bn function: this can only contain unions or
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   165
   appends of elements; in case of recursive calls it retruns also the applied
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   166
   bn function *)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
fun strip_bn_fun t =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  case t of
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
    Const (@{const_name sup}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  | Const (@{const_name append}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  | Const (@{const_name insert}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
      (i, NONE) :: strip_bn_fun y
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  | Const (@{const_name Cons}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
      (i, NONE) :: strip_bn_fun y
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  | Const (@{const_name bot}, _) => []
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  | Const (@{const_name Nil}, _) => []
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  | (f as Free _) $ Bound i => [(i, SOME f)]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
  | _ => error ("Unsupported binding function: " ^ (PolyML.makestring t))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
fun find [] _ = error ("cannot find element")
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  | find ((x, z)::xs) y = if (Long_Name.base_name x) = y then z else find xs y
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
ML {*
2122
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   188
fun prep_bn lthy dt_names dts eqs = 
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  fun aux eq = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
    val (lhs, rhs) = eq
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
      |> strip_qnt_body "all" 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
      |> HOLogic.dest_Trueprop
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
      |> HOLogic.dest_eq
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
    val (bn_fun, [cnstr]) = strip_comb lhs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
    val (_, ty) = dest_Free bn_fun
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
    val (ty_name, _) = dest_Type (domain_type ty)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
    val dt_index = find_index (fn x => x = ty_name) dt_names
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
    val (cnstr_head, cnstr_args) = strip_comb cnstr
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
    val rhs_elements = strip_bn_fun rhs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
    val included = map (apfst (fn i => length (cnstr_args) - i - 1)) rhs_elements
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
    (dt_index, (bn_fun, (cnstr_head, included)))
2122
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   205
  end
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  fun order dts i ts = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
    val dt = nth dts i
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
    val cts = map (fn (x, _, _) => Binding.name_of x) ((fn (_, _, _, x) => x) dt)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
    val ts' = map (fn (x, y) => (fst (dest_Const x), y)) ts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
    map (find ts') cts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  val unordered = AList.group (op=) (map aux eqs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  val unordered' = map (fn (x, y) =>  (x, AList.group (op=) y)) unordered
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
  val ordered = map (fn (x, y) => (x, map (fn (v, z) => (v, order dts x z)) y)) unordered' 
2118
0e52851acac4 moved the data-transformation into the parser
Christian Urban <urbanc@in.tum.de>
parents: 2107
diff changeset
   218
  val ordered' = flat (map (fn (ith, l) => map (fn (bn, data) => (bn, ith, data)) l) ordered)
2122
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   219
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   220
  (*val _ = tracing ("eqs\n" ^ cat_lines (map (Syntax.string_of_term lthy) eqs))*)
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   221
  (*val _ = tracing ("map eqs\n" ^ @{make_string} (map aux2 eqs))*)
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   222
  (*val _ = tracing ("ordered'\n" ^ @{make_string} ordered')*)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
in
2125
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   224
  ordered'
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
fun raw_nominal_decls dts bn_funs bn_eqs binds lthy =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
  val thy = ProofContext.theory_of lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  val thy_name = Context.theory_name thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  val dt_names = map (fn (_, s, _, _) => Binding.name_of s) dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  val dt_full_names = map (Long_Name.qualify thy_name) dt_names 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  val dt_full_names' = add_raws dt_full_names
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
  val dts_env = dt_full_names ~~ dt_full_names'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  val cnstrs = get_cnstr_strs dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
  val cnstrs_ty = get_typed_cnstrs dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
  val cnstrs_full_names = map (Long_Name.qualify thy_name) cnstrs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  val cnstrs_full_names' = map (fn (x, y) => Long_Name.qualify thy_name 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
    (Long_Name.qualify (add_raw x) (add_raw y))) cnstrs_ty
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  val cnstrs_env = cnstrs_full_names ~~ cnstrs_full_names'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
  val bn_fun_strs = get_bn_fun_strs bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  val bn_fun_strs' = add_raws bn_fun_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  val bn_fun_env = bn_fun_strs ~~ bn_fun_strs'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  val bn_fun_full_env = map (pairself (Long_Name.qualify thy_name)) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
    (bn_fun_strs ~~ bn_fun_strs')
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  val (raw_dt_names, raw_dts) = rawify_dts dt_names dts dts_env
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  val raw_bclauses = rawify_bclauses dts_env cnstrs_env bn_fun_full_env binds 
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   255
  val raw_bns = prep_bn lthy dt_full_names' raw_dts (map snd raw_bn_eqs)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   257
  val (raw_dt_names, lthy1) = add_datatype_wrapper raw_dt_names raw_dts lthy
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   258
  val (raw_bn_funs, raw_bn_eqs, lthy2) = add_primrec_wrapper raw_bn_funs raw_bn_eqs lthy1
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   259
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   260
  val morphism_2_0 = ProofContext.export_morphism lthy2 lthy
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   261
  fun export_fun f (t, n , l) = (f t, n, map (map (apsnd (Option.map f))) l);
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   262
  val bn_funs_decls = map (export_fun (Morphism.term morphism_2_0)) raw_bns;
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
in
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   264
  (raw_dt_names, raw_bn_eqs, raw_bclauses, bn_funs_decls, raw_bns, lthy2)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
2017
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   268
lemma equivp_hack: "equivp x"
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   269
sorry
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   270
ML {*
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   271
fun equivp_hack ctxt rel =
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   272
let
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   273
  val thy = ProofContext.theory_of ctxt
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   274
  val ty = domain_type (fastype_of rel)
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   275
  val cty = ctyp_of thy ty
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   276
  val ct = cterm_of thy rel
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   277
in
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   278
  Drule.instantiate' [SOME cty] [SOME ct] @{thm equivp_hack}
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   279
end
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   280
*}
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   281
2023
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   282
ML {* val cheat_equivp = Unsynchronized.ref false *}
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   283
ML {* val cheat_fv_rsp = Unsynchronized.ref false *}
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   284
ML {* val cheat_alpha_bn_rsp = Unsynchronized.ref false *}
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   285
ML {* val cheat_supp_eq = Unsynchronized.ref false *}
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   286
2025
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   287
ML {*
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   288
fun remove_loop t =
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   289
  let val _ = HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of t)) in t end
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   290
  handle TERM _ => @{thm eqTrueI} OF [t]
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   291
*}
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   292
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   293
text {* 
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   294
  nominal_datatype2 does the following things in order:
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   295
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   296
Parser.thy/raw_nominal_decls
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   297
  1) define the raw datatype
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   298
  2) define the raw binding functions 
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   299
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   300
Perm.thy/define_raw_perms
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   301
  3) define permutations of the raw datatype and show that the raw type is 
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   302
     in the pt typeclass
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   303
      
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   304
Lift.thy/define_fv_alpha_export, Fv.thy/define_fv & define_alpha
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   305
  4) define fv and fv_bn
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   306
  5) define alpha and alpha_bn
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   307
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   308
Perm.thy/distinct_rel
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   309
  6) prove alpha_distincts (C1 x \<notsimeq> C2 y ...)             (Proof by cases; simp)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   310
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   311
Tacs.thy/build_rel_inj
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   312
  6) prove alpha_eq_iff    (C1 x = C2 y \<leftrightarrow> P x y ...)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   313
     (left-to-right by intro rule, right-to-left by cases; simp)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   314
Equivp.thy/prove_eqvt
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   315
  7) prove bn_eqvt (common induction on the raw datatype)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   316
  8) prove fv_eqvt (common induction on the raw datatype with help of above)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   317
Rsp.thy/build_alpha_eqvts
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   318
  9) prove alpha_eqvt and alpha_bn_eqvt
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   319
     (common alpha-induction, unfolding alpha_gen, permute of #* and =)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   320
Equivp.thy/build_alpha_refl & Equivp.thy/build_equivps
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   321
 10) prove that alpha and alpha_bn are equivalence relations
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   322
     (common induction and application of 'compose' lemmas)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   323
Lift.thy/define_quotient_types
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   324
 11) define quotient types
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   325
Rsp.thy/build_fvbv_rsps
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   326
 12) prove bn respects     (common induction and simp with alpha_gen)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   327
Rsp.thy/prove_const_rsp
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   328
 13) prove fv respects     (common induction and simp with alpha_gen)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   329
 14) prove permute respects    (unfolds to alpha_eqvt)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   330
Rsp.thy/prove_alpha_bn_rsp
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   331
 15) prove alpha_bn respects
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   332
     (alpha_induct then cases then sym and trans of the relations)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   333
Rsp.thy/prove_alpha_alphabn
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   334
 16) show that alpha implies alpha_bn (by unduction, needed in following step)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   335
Rsp.thy/prove_const_rsp
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   336
 17) prove respects for all datatype constructors
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   337
     (unfold eq_iff and alpha_gen; introduce zero permutations; simp)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   338
Perm.thy/quotient_lift_consts_export
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   339
 18) define lifted constructors, fv, bn, alpha_bn, permutations
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   340
Perm.thy/define_lifted_perms
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   341
 19) lift permutation zero and add properties to show that quotient type is in the pt typeclass
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   342
Lift.thy/lift_thm
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   343
 20) lift permutation simplifications
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   344
 21) lift induction
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   345
 22) lift fv
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   346
 23) lift bn
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   347
 24) lift eq_iff
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   348
 25) lift alpha_distincts
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   349
 26) lift fv and bn eqvts
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   350
Equivp.thy/prove_supports
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   351
 27) prove that union of arguments supports constructors
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   352
Equivp.thy/prove_fs
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   353
 28) show that the lifted type is in fs typeclass     (* by q_induct, supports *)
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   354
Equivp.thy/supp_eq
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   355
 29) prove supp = fv
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   356
*}
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   357
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: 2037
diff changeset
   358
ML {*
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: 2037
diff changeset
   359
(* for testing porposes - to exit the procedure early *)
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: 2037
diff changeset
   360
exception TEST of Proof.context
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: 2037
diff changeset
   361
2125
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   362
val (STEPS, STEPS_setup) = Attrib.config_int "STEPS" (K 10);
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   363
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   364
fun get_STEPS ctxt = Config.get ctxt STEPS
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: 2037
diff changeset
   365
*}
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   366
2125
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   367
setup STEPS_setup
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   368
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
fun nominal_datatype2 dts bn_funs bn_eqs bclauses lthy =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
let
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   372
  (* definition of the raw datatypes and raw bn-functions *)
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   373
  val (raw_dt_names, raw_bn_eqs, raw_bclauses, raw_bns, raw_bns2, lthy1) =
2125
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   374
    if get_STEPS lthy > 1 then raw_nominal_decls dts bn_funs bn_eqs bclauses lthy
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: 2037
diff changeset
   375
    else raise TEST lthy
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   377
  (*val _ = tracing ("exported: " ^ commas (map @{make_string} raw_bns))*)
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   378
  (*val _ = tracing ("plain: " ^ commas (map @{make_string} raw_bns2))*)
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   379
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   380
  val dtinfo = Datatype.the_info (ProofContext.theory_of lthy1) (hd raw_dt_names)
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   381
  val {descr, sorts, ...} = dtinfo
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   382
  val all_tys = map (fn (i, _) => nth_dtyp descr sorts i) descr
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   383
  val all_full_tnames = map (fn (_, (n, _, _)) => n) descr
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   384
  val dtinfos = map (Datatype.the_info (ProofContext.theory_of lthy1)) all_full_tnames
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: 2037
diff changeset
   385
  
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: 2037
diff changeset
   386
  val inject_thms = flat (map #inject dtinfos);
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: 2037
diff changeset
   387
  val distinct_thms = flat (map #distinct dtinfos);
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: 2037
diff changeset
   388
  val rel_dtinfos = List.take (dtinfos, (length dts)); 
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: 2037
diff changeset
   389
  val rel_distinct = map #distinct rel_dtinfos;  (* thm list list *)
2047
31ba33a199c7 fixed my error with define_raw_fv
Christian Urban <urbanc@in.tum.de>
parents: 2046
diff changeset
   390
  val induct_thm = #induct dtinfo;
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: 2037
diff changeset
   391
  val exhaust_thms = map #exhaust dtinfos;
1999
4df3441aa0b4 more parser/new parser synchronization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   392
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   393
  (* definitions of raw permutations *)
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   394
  val ((perms, raw_perm_defs, raw_perm_simps), lthy2) =
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   395
    if get_STEPS lthy1 > 2 
2125
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   396
    then Local_Theory.theory_result (define_raw_perms descr sorts induct_thm (length dts)) lthy1
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: 2037
diff changeset
   397
    else raise TEST lthy1
2144
e900526e95c4 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 2143
diff changeset
   398
 
e900526e95c4 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 2143
diff changeset
   399
  (* noting the raw permutations as eqvt theorems *)
e900526e95c4 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 2143
diff changeset
   400
  val eqvt_attrib = Attrib.internal (K Nominal_ThmDecls.eqvt_add)
e900526e95c4 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 2143
diff changeset
   401
  val (_, lthy2a) = Local_Theory.note ((Binding.empty, [eqvt_attrib]), raw_perm_defs) lthy2
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   402
2105
e25b0fff0dd2 Include raw permutation definitions in eqvt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2089
diff changeset
   403
  val thy = Local_Theory.exit_global lthy2a;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   404
  val thy_name = Context.theory_name thy
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   405
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   406
  (* definition of raw fv_functions *)
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1945
diff changeset
   407
  val lthy3 = Theory_Target.init NONE thy;
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   408
  val raw_bn_funs = map (fn (f, _, _) => f) raw_bns;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1945
diff changeset
   409
2146
a2f70c09e77d minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2144
diff changeset
   410
  val (fv, fvbn, fv_def, lthy3a) = 
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   411
    if get_STEPS lthy2 > 3 
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   412
    then define_raw_fvs descr sorts raw_bns raw_bns2 raw_bclauses lthy3
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: 2037
diff changeset
   413
    else raise TEST lthy3
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   414
  
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   415
  (* definition of raw alphas *)
1999
4df3441aa0b4 more parser/new parser synchronization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   416
  val (((alpha_ts, alpha_intros), (alpha_cases, alpha_induct)), lthy4) =
2125
60ee289a8c63 made out of STEPS a configuration value so that it can be set individually in each file
Christian Urban <urbanc@in.tum.de>
parents: 2122
diff changeset
   417
    if get_STEPS lthy > 4 
2142
c39d4fe31100 moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   418
    then define_raw_alpha dtinfo raw_bns raw_bclauses fv lthy3a
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: 2037
diff changeset
   419
    else raise TEST lthy3a
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: 2037
diff changeset
   420
  
1999
4df3441aa0b4 more parser/new parser synchronization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   421
  val (alpha_ts_nobn, alpha_ts_bn) = chop (length fv) alpha_ts
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   422
  
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   423
  val dts_names = map (fn (i, (s, _, _)) => (s, i)) (#descr dtinfo);
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   424
  val bn_tys = map (domain_type o fastype_of) raw_bn_funs;
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   425
  val bn_nos = map (dtyp_no_of_typ dts_names) bn_tys;
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   426
  val bns = raw_bn_funs ~~ bn_nos;
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   427
  val rel_dists = flat (map (distinct_rel lthy4 alpha_cases)
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   428
    (rel_distinct ~~ alpha_ts_nobn));
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   429
  val rel_dists_bn = flat (map (distinct_rel lthy4 alpha_cases)
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   430
    ((map (fn i => nth rel_distinct i) bn_nos) ~~ alpha_ts_bn))
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   431
  
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   432
  (* definition of raw_alpha_eq_iff  lemmas *)
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: 2037
diff changeset
   433
  val alpha_eq_iff = build_rel_inj alpha_intros (inject_thms @ distinct_thms) alpha_cases lthy4
2025
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   434
  val alpha_eq_iff_simp = map remove_loop alpha_eq_iff;
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   435
  
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   436
  (*val _ = map tracing (map PolyML.makestring alpha_eq_iff_simp);*)
2022
8ffede2c8ce9 Introduce eq_iff_simp to match the one from Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2020
diff changeset
   437
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   438
  (* proving equivariance lemmas *)
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   439
  val _ = warning "Proving equivariance";
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   440
  val (bv_eqvt, lthy5) = prove_eqvt all_tys induct_thm (raw_bn_eqs @ raw_perm_defs) (map fst bns) lthy4
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   441
  val (fv_eqvt, lthy6) = prove_eqvt all_tys induct_thm (fv_def @ raw_perm_defs) (fv @ fvbn) lthy5
2114
3201a8c3456b Use equivariance instead of alpha_eqvt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2107
diff changeset
   442
  val (alpha_eqvt, lthy6a) = Nominal_Eqvt.equivariance alpha_ts alpha_induct alpha_intros lthy6;
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   443
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: 2037
diff changeset
   444
  (* proving alpha equivalence *)
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   445
  val _ = warning "Proving equivalence";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   446
  val fv_alpha_all = combine_fv_alpha_bns (fv, fvbn) (alpha_ts_nobn, alpha_ts_bn) bn_nos;
2114
3201a8c3456b Use equivariance instead of alpha_eqvt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2107
diff changeset
   447
  val reflps = build_alpha_refl fv_alpha_all alpha_ts induct_thm alpha_eq_iff_simp lthy6a;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   448
  val alpha_equivp =
2114
3201a8c3456b Use equivariance instead of alpha_eqvt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2107
diff changeset
   449
    if !cheat_equivp then map (equivp_hack lthy6a) alpha_ts
2017
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   450
    else build_equivps alpha_ts reflps alpha_induct
2114
3201a8c3456b Use equivariance instead of alpha_eqvt
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2107
diff changeset
   451
      inject_thms alpha_eq_iff_simp distinct_thms alpha_cases alpha_eqvt lthy6a;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   452
  val qty_binds = map (fn (_, b, _, _) => b) dts;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   453
  val qty_names = map Name.of_binding qty_binds;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   454
  val qty_full_names = map (Long_Name.qualify thy_name) qty_names
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   455
  val (qtys, lthy7) = define_quotient_types qty_binds all_tys alpha_ts_nobn alpha_equivp lthy6a;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   456
  val const_names = map Name.of_binding (flat (map (fn (_, _, _, t) => map (fn (b, _, _) => b) t) dts));
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   457
  val raw_consts =
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   458
    flat (map (fn (i, (_, _, l)) =>
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   459
      map (fn (cname, dts) =>
2016
1715dfe9406c Fix Datatype_Aux calls in NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2011
diff changeset
   460
        Const (cname, map (Datatype_Aux.typ_of_dtyp descr sorts) dts --->
1715dfe9406c Fix Datatype_Aux calls in NewParser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2011
diff changeset
   461
          Datatype_Aux.typ_of_dtyp descr sorts (Datatype_Aux.DtRec i))) l) descr);
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   462
  val (consts, const_defs, lthy8) = quotient_lift_consts_export qtys (const_names ~~ raw_consts) lthy7;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   463
  val _ = warning "Proving respects";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   464
  val bns_rsp_pre' = build_fvbv_rsps alpha_ts alpha_induct raw_bn_eqs (map fst bns) lthy8;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   465
  val (bns_rsp_pre, lthy9) = fold_map (
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   466
    fn (bn_t, _) => prove_const_rsp qtys Binding.empty [bn_t] (fn _ =>
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   467
       resolve_tac bns_rsp_pre' 1)) bns lthy8;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   468
  val bns_rsp = flat (map snd bns_rsp_pre);
2017
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   469
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   470
  fun fv_rsp_tac _ = if !cheat_fv_rsp then Skip_Proof.cheat_tac thy
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   471
    else fvbv_rsp_tac alpha_induct fv_def lthy8 1;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   472
  val fv_rsps = prove_fv_rsp fv_alpha_all alpha_ts fv_rsp_tac lthy9;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   473
  val (fv_rsp_pre, lthy10) = fold_map
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   474
    (fn fv => fn ctxt => prove_const_rsp qtys Binding.empty [fv]
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   475
    (fn _ => asm_simp_tac (HOL_ss addsimps fv_rsps) 1) ctxt) (fv @ fvbn) lthy9;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   476
  val fv_rsp = flat (map snd fv_rsp_pre);
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   477
  val (perms_rsp, lthy11) = prove_const_rsp qtys Binding.empty perms
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   478
    (fn _ => asm_simp_tac (HOL_ss addsimps alpha_eqvt) 1) lthy10;
2023
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   479
  fun alpha_bn_rsp_tac _ = if !cheat_alpha_bn_rsp then Skip_Proof.cheat_tac thy
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   480
    else
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: 2037
diff changeset
   481
      let val alpha_bn_rsp_pre = prove_alpha_bn_rsp alpha_ts alpha_induct (alpha_eq_iff_simp @ rel_dists @ rel_dists_bn) alpha_equivp exhaust_thms alpha_ts_bn lthy11 in asm_simp_tac (HOL_ss addsimps alpha_bn_rsp_pre) 1 end;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   482
  val (alpha_bn_rsps, lthy11a) = fold_map (fn cnst => prove_const_rsp qtys Binding.empty [cnst]
2023
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   483
    alpha_bn_rsp_tac) alpha_ts_bn lthy11
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   484
  fun const_rsp_tac _ =
2022
8ffede2c8ce9 Introduce eq_iff_simp to match the one from Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2020
diff changeset
   485
    let val alpha_alphabn = prove_alpha_alphabn alpha_ts alpha_induct alpha_eq_iff_simp alpha_ts_bn lthy11a
8ffede2c8ce9 Introduce eq_iff_simp to match the one from Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2020
diff changeset
   486
      in constr_rsp_tac alpha_eq_iff_simp (fv_rsp @ bns_rsp @ reflps @ alpha_alphabn) 1 end
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   487
  val (const_rsps, lthy12) = fold_map (fn cnst => prove_const_rsp qtys Binding.empty [cnst]
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   488
    const_rsp_tac) raw_consts lthy11a
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   489
    val qfv_names = map (unsuffix "_raw" o Long_Name.base_name o fst o dest_Const) (fv @ fvbn)
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   490
  val (qfv_ts, qfv_defs, lthy12a) = quotient_lift_consts_export qtys (qfv_names ~~ (fv @ fvbn)) lthy12;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   491
  val (qfv_ts_nobn, qfv_ts_bn) = chop (length perms) qfv_ts;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   492
  val qbn_names = map (fn (b, _ , _) => Name.of_binding b) bn_funs
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   493
  val (qbn_ts, qbn_defs, lthy12b) = quotient_lift_consts_export qtys (qbn_names ~~ raw_bn_funs) lthy12a;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   494
  val qalpha_bn_names = map (unsuffix "_raw" o Long_Name.base_name o fst o dest_Const) alpha_ts_bn
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   495
  val (qalpha_ts_bn, qalphabn_defs, lthy12c) = 
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   496
    quotient_lift_consts_export qtys (qalpha_bn_names ~~ alpha_ts_bn) lthy12b;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   497
  val _ = warning "Lifting permutations";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   498
  val thy = Local_Theory.exit_global lthy12c;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   499
  val perm_names = map (fn x => "permute_" ^ x) qty_names
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   500
  val thy' = define_lifted_perms qtys qty_full_names (perm_names ~~ perms) raw_perm_simps thy;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   501
  val lthy13 = Theory_Target.init NONE thy';
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   502
  val q_name = space_implode "_" qty_names;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   503
  fun suffix_bind s = Binding.qualify true q_name (Binding.name s);
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   504
  val _ = warning "Lifting induction";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   505
  val constr_names = map (Long_Name.base_name o fst o dest_Const) consts;
2047
31ba33a199c7 fixed my error with define_raw_fv
Christian Urban <urbanc@in.tum.de>
parents: 2046
diff changeset
   506
  val q_induct = Rule_Cases.name constr_names (lift_thm qtys lthy13 induct_thm);
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   507
  fun note_suffix s th ctxt =
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   508
    snd (Local_Theory.note ((suffix_bind s, []), th) ctxt);
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   509
  fun note_simp_suffix s th ctxt =
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   510
    snd (Local_Theory.note ((suffix_bind s, [Attrib.internal (K Simplifier.simp_add)]), th) ctxt);
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   511
  val (_, lthy14) = Local_Theory.note ((suffix_bind "induct",
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   512
    [Attrib.internal (K (Rule_Cases.case_names constr_names))]), 
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   513
    [Rule_Cases.name constr_names q_induct]) lthy13;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   514
  val q_inducts = Project_Rule.projects lthy13 (1 upto (length fv)) q_induct
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   515
  val (_, lthy14a) = Local_Theory.note ((suffix_bind "inducts", []), q_inducts) lthy14;
2143
871d8a5e0c67 somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
Christian Urban <urbanc@in.tum.de>
parents: 2142
diff changeset
   516
  val q_perm = map (lift_thm qtys lthy14) raw_perm_defs;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   517
  val lthy15 = note_simp_suffix "perm" q_perm lthy14a;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   518
  val q_fv = map (lift_thm qtys lthy15) fv_def;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   519
  val lthy16 = note_simp_suffix "fv" q_fv lthy15;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   520
  val q_bn = map (lift_thm qtys lthy16) raw_bn_eqs;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   521
  val lthy17 = note_simp_suffix "bn" q_bn lthy16;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   522
  val _ = warning "Lifting eq-iff";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   523
  (*val _ = map tracing (map PolyML.makestring alpha_eq_iff);*)
2089
e9f089a5cc17 Parser changes for compound relations
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2076
diff changeset
   524
  val eq_iff_unfolded0 = map (Local_Defs.unfold lthy17 @{thms alphas}) alpha_eq_iff_simp
e9f089a5cc17 Parser changes for compound relations
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2076
diff changeset
   525
  val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms Pair_eqvt}) eq_iff_unfolded0
e9f089a5cc17 Parser changes for compound relations
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2076
diff changeset
   526
  val q_eq_iff_pre0 = map (lift_thm qtys lthy17) eq_iff_unfolded1;
e9f089a5cc17 Parser changes for compound relations
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2076
diff changeset
   527
  val q_eq_iff_pre1 = map (Local_Defs.fold lthy17 @{thms Pair_eqvt}) q_eq_iff_pre0
e9f089a5cc17 Parser changes for compound relations
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2076
diff changeset
   528
  val q_eq_iff_pre2 = map (Local_Defs.fold lthy17 @{thms alphas}) q_eq_iff_pre1
e9f089a5cc17 Parser changes for compound relations
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2076
diff changeset
   529
  val q_eq_iff = map (Local_Defs.unfold lthy17 (Quotient_Info.id_simps_get lthy17)) q_eq_iff_pre2
2025
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   530
  val (_, lthy18) = Local_Theory.note ((suffix_bind "eq_iff", []), q_eq_iff) lthy17;
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   531
  val q_dis = map (lift_thm qtys lthy18) rel_dists;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   532
  val lthy19 = note_simp_suffix "distinct" q_dis lthy18;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   533
  val q_eqvt = map (lift_thm qtys lthy19) (bv_eqvt @ fv_eqvt);
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   534
  val (_, lthy20) = Local_Theory.note ((Binding.empty,
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   535
    [Attrib.internal (fn _ => Nominal_ThmDecls.eqvt_add)]), q_eqvt) lthy19;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   536
  val _ = warning "Supports";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   537
  val supports = map (prove_supports lthy20 q_perm) consts;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   538
  val fin_supp = HOLogic.conj_elims (prove_fs lthy20 q_induct supports qtys);
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   539
  val thy3 = Local_Theory.exit_global lthy20;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   540
  val _ = warning "Instantiating FS";
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   541
  val lthy21 = Theory_Target.instantiation (qty_full_names, [], @{sort fs}) thy3;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   542
  fun tac _ = Class.intro_classes_tac [] THEN (ALLGOALS (resolve_tac fin_supp))
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   543
  val lthy22 = Class.prove_instantiation_instance tac lthy21
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   544
  val fv_alpha_all = combine_fv_alpha_bns (qfv_ts_nobn, qfv_ts_bn) (alpha_ts_nobn, qalpha_ts_bn) bn_nos;
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   545
  val (names, supp_eq_t) = supp_eq fv_alpha_all;
2020
8468be06bff1 Cheat support equations in new parser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2017
diff changeset
   546
  val _ = warning "Support Equations";
2023
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   547
  fun supp_eq_tac' _ = if !cheat_supp_eq then Skip_Proof.cheat_tac thy else
2025
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   548
    supp_eq_tac q_induct q_fv q_perm q_eq_iff lthy22 1;
2023
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   549
  val q_supp = HOLogic.conj_elims (Goal.prove lthy22 names [] supp_eq_t supp_eq_tac') handle e =>
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   550
    let val _ = warning ("Support eqs failed") in [] end;
2020
8468be06bff1 Cheat support equations in new parser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2017
diff changeset
   551
  val lthy23 = note_suffix "supp" q_supp lthy22;
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
in
2020
8468be06bff1 Cheat support equations in new parser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2017
diff changeset
   553
  (0, lthy23)
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: 2037
diff changeset
   554
end handle TEST ctxt => (0, ctxt)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
section {* Preparing and parsing of the specification *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
(* parsing the datatypes and declaring *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
(* constructors in the local theory    *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
fun prepare_dts dt_strs lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
  val thy = ProofContext.theory_of lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
  fun mk_type full_tname tvrs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
    Type (full_tname, map (fn a => TVar ((a, 0), [])) tvrs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
  fun prep_cnstr full_tname tvs (cname, anno_tys, mx, _) =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
    val tys = map (Syntax.read_typ lthy o snd) anno_tys
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
    val ty = mk_type full_tname tvs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
    ((cname, tys ---> ty, mx), (cname, tys, mx))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
  fun prep_dt (tvs, tname, mx, cnstrs) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
    val full_tname = Sign.full_name thy tname
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
    val (cnstrs', cnstrs'') = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
      split_list (map (prep_cnstr full_tname tvs) cnstrs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
    (cnstrs', (tvs, tname, mx, cnstrs''))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
  end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
  val (cnstrs, dts) = split_list (map prep_dt dt_strs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
  lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
  |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
  |> pair dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
(* parsing the binding function specification and *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
(* declaring the functions in the local theory    *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
  val ((bn_funs, bn_eqs), _) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
    Specification.read_spec bn_fun_strs bn_eq_strs lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
  fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  val bn_funs' = map prep_bn_fun bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
  lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
  |> Local_Theory.theory (Sign.add_consts_i bn_funs')
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
  |> pair (bn_funs', bn_eqs) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
text {* associates every SOME with the index in the list; drops NONEs *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
fun indexify xs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
  fun mapp _ [] = []
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
    | mapp i (NONE :: xs) = mapp (i + 1) xs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
    | mapp i (SOME x :: xs) = (x, i) :: mapp (i + 1) xs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
in 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
  mapp 0 xs 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
fun index_lookup xs x =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
  case AList.lookup (op=) xs x of
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
    SOME x => x
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
  | NONE => error ("Cannot find " ^ x ^ " as argument annotation.");
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
fun prepare_bclauses dt_strs lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
  val annos_bclauses =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
    get_cnstrs dt_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
    |> map (map (fn (_, antys, _, bns) => (map fst antys, bns)))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
  fun prep_binder env bn_str =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
    case (Syntax.read_term lthy bn_str) of
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
      Free (x, _) => (NONE, index_lookup env x)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
    | Const (a, T) $ Free (x, _) => (SOME (Const (a, T)), index_lookup env x)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
    | _ => error ("The term " ^ bn_str ^ " is not allowed as binding function.")
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
  fun prep_body env bn_str = index_lookup env bn_str
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   644
  fun prep_mode "bind"     = Lst 
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   645
    | prep_mode "bind_set" = Set 
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   646
    | prep_mode "bind_res" = Res 
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
  fun prep_bclause env (mode, binders, bodies) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
    val binders' = map (prep_binder env) binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
    val bodies' = map (prep_body env) bodies
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
  in  
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   653
    BC (prep_mode mode, binders', bodies')
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
  fun prep_bclauses (annos, bclause_strs) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
    val env = indexify annos (* for every label, associate the index *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
    map (prep_bclause env) bclause_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  map (map prep_bclauses) annos_bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   667
text {* 
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   668
  adds an empty binding clause for every argument
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   669
  that is not already part of a binding clause
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   670
*}
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   671
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
fun included i bcs = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
let
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   675
  fun incl (BC (_, bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
  exists incl bcs 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   681
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   682
fun complete dt_strs bclauses = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
  val args = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
    get_cnstrs dt_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
    |> map (map (fn (_, antys, _, _) => length antys))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
  fun complt n bcs = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
  let
2288
3b83960f9544 new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents: 2146
diff changeset
   690
    fun add bcs i = (if included i bcs then [] else [BC (Lst, [], [i])]) 
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
    bcs @ (flat (map_range (add bcs) n))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
  map2 (map2 complt) args bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
  fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
  val lthy0 = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
    Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) lthy
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
   705
  val (dts, lthy1) = prepare_dts dt_strs lthy0
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
   706
  val ((bn_funs, bn_eqs), lthy2) = prepare_bn_funs bn_fun_strs bn_eq_strs lthy1
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
  val bclauses = prepare_bclauses dt_strs lthy2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
  val bclauses' = complete dt_strs bclauses 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
  nominal_datatype2 dts bn_funs bn_eqs bclauses' lthy |> snd
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
(* Command Keyword *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
val _ = OuterSyntax.local_theory "nominal_datatype" "test" OuterKeyword.thy_decl
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
  (main_parser >> nominal_datatype2_cmd)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
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: 2037
diff changeset
   720
(*
2035
3622cae9b10e to my best knowledge the number of datatypes is equal to the length of the dt_descr; so we can save one argument in define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents: 2027
diff changeset
   721
atom_decl name
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
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
   723
nominal_datatype lam =
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   724
  Var name
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   725
| App lam lam
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   726
| Lam x::name t::lam  bind_set x in t
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   727
| Let p::pt t::lam bind_set "bn p" in t
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   728
and pt =
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   729
  P1 name
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   730
| P2 pt pt
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   731
binder
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   732
 bn::"pt \<Rightarrow> atom set"
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
   733
where
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   734
  "bn (P1 x) = {atom x}"
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   735
| "bn (P2 p1 p2) = bn p1 \<union> bn p2"
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   736
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: 2037
diff changeset
   737
find_theorems Var_raw
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: 2037
diff changeset
   738
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: 2037
diff changeset
   739
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: 2037
diff changeset
   740
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   741
thm lam_pt.bn
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   742
thm lam_pt.fv[simplified lam_pt.supp(1-2)]
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   743
thm lam_pt.eq_iff
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   744
thm lam_pt.induct
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   745
thm lam_pt.perm
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
nominal_datatype exp =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
  EVar name
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
| EUnit
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   750
| EPair q1::exp q2::exp
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
| ELetRec l::lrbs e::exp     bind "b_lrbs l" in e l
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
and fnclause =
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   754
  K x::name p::pat f::exp    bind_res "b_pat p" in f
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
and fnclauses =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
  S fnclause
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
| ORs fnclause fnclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   759
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
and lrb =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
  Clause fnclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
and lrbs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   764
  Single lrb
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   765
| More lrb lrbs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
and pat =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   768
  PVar name
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
| PUnit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
| PPair pat pat
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
binder
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   773
  b_lrbs      :: "lrbs \<Rightarrow> atom list" and
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
  b_pat       :: "pat \<Rightarrow> atom set" and
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   775
  b_fnclauses :: "fnclauses \<Rightarrow> atom list" and
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   776
  b_fnclause  :: "fnclause \<Rightarrow> atom list" and
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   777
  b_lrb       :: "lrb \<Rightarrow> atom list"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
where
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
  "b_lrbs (Single l) = b_lrb l"
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   781
| "b_lrbs (More l ls) = append (b_lrb l) (b_lrbs ls)"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
| "b_pat (PVar x) = {atom x}"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
| "b_pat (PUnit) = {}"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
| "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
| "b_fnclauses (S fc) = (b_fnclause fc)"
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   786
| "b_fnclauses (ORs fc fcs) = append (b_fnclause fc) (b_fnclauses fcs)"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
| "b_lrb (Clause fcs) = (b_fnclauses fcs)"
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   788
| "b_fnclause (K x pat exp) = [atom x]"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   790
thm exp_fnclause_fnclauses_lrb_lrbs_pat.bn
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   791
thm exp_fnclause_fnclauses_lrb_lrbs_pat.fv
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   792
thm exp_fnclause_fnclauses_lrb_lrbs_pat.eq_iff
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   793
thm exp_fnclause_fnclauses_lrb_lrbs_pat.induct
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   794
thm exp_fnclause_fnclauses_lrb_lrbs_pat.perm
1945
93e5a31ba230 added definition of raw-permutations to the new-parser
Christian Urban <urbanc@in.tum.de>
parents: 1944
diff changeset
   795
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   796
nominal_datatype ty =
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   797
  Vr "name"
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   798
| Fn "ty" "ty"
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   799
and tys =
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   800
  Al xs::"name fset" t::"ty" bind_res xs in t
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   801
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   802
thm ty_tys.fv[simplified ty_tys.supp]
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   803
thm ty_tys.eq_iff
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   804
2017
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   805
*)
2001
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   806
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   807
(* some further tests *)
7c8242a02f39 NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2000
diff changeset
   808
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   809
(*
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   810
nominal_datatype ty2 =
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   811
  Vr2 "name"
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   812
| Fn2 "ty2" "ty2"
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   813
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   814
nominal_datatype tys2 =
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   815
  All2 xs::"name fset" ty::"ty2" bind_res xs in ty
1964
209ee65b2395 added some further problemetic tests
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   816
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
   817
nominal_datatype lam2 =
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   818
  Var2 "name"
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   819
| App2 "lam2" "lam2 list"
8daf6ff5e11a simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1970
diff changeset
   820
| Lam2 x::"name" t::"lam2" bind x in t
2008
1bddffddc03f attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
Christian Urban <urbanc@in.tum.de>
parents: 2007
diff changeset
   821
*)
1964
209ee65b2395 added some further problemetic tests
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   822
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828