Nominal/Parser.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Sat, 12 Jun 2010 06:35:27 +0200
changeset 2225 cbffed7d81bd
parent 2138 2e514a0aca63
permissions -rw-r--r--
Fix integer relation.
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" 
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
     5
        "Perm" "NewFv" "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
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
section{* Interface for nominal_datatype *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
(* nominal datatype parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
local
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  structure P = OuterParse;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  structure S = Scan
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  fun triple1 ((x, y), z) = (x, y, z)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  fun triple2 (x, (y, z)) = (x, y, z)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  fun tuple ((x, y, z), u) = (x, y, z, u)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  fun tswap (((x, y), z), u) = (x, y, u, z)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
val _ = OuterKeyword.keyword "bind"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
val _ = OuterKeyword.keyword "bind_set"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
val _ = OuterKeyword.keyword "bind_res"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
    27
val anno_typ = S.option (P.name --| P.$$$ "::") -- P.typ
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
val bind_mode = P.$$$ "bind" || P.$$$ "bind_set" || P.$$$ "bind_res"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
val bind_clauses = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  P.enum "," (bind_mode -- S.repeat1 P.term -- (P.$$$ "in" |-- S.repeat1 P.name) >> triple1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
val cnstr_parser =
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
    35
  P.binding -- S.repeat anno_typ -- bind_clauses -- P.opt_mixfix >> tswap
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
(* datatype parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
val dt_parser =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  (P.type_args -- P.binding -- P.opt_mixfix >> triple1) -- 
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
    40
    (P.$$$ "=" |-- P.enum1 "|" cnstr_parser) >> tuple
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
(* binding function parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
val bnfun_parser = 
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    44
  S.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([], [])
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
(* main parser *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
val main_parser =
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    48
  P.and_list1 dt_parser -- bnfun_parser >> triple2
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
fun get_cnstrs dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  map (fn (_, _, _, constrs) => constrs) dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
fun get_typed_cnstrs dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  flat (map (fn (_, bn, _, constrs) => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
   (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
fun get_cnstr_strs dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  map (fn (bn, _, _) => Binding.name_of bn) (flat (get_cnstrs dts))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
fun get_bn_fun_strs bn_funs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
    68
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
fun add_primrec_wrapper funs eqs lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  if null funs then (([], []), lthy)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  else 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
   let 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
     val eqs' = map (fn (_, eq) => (Attrib.empty_binding, eq)) eqs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
     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
    76
     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
    77
     val phi = ProofContext.export_morphism lthy' lthy
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
   in 
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
    79
     ((map (Morphism.term phi) funs'', map (Morphism.thm phi) eqs''), lthy')
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
   end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
fun add_datatype_wrapper dt_names dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  val conf = Datatype.default_config
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    92
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    93
text {* Infrastructure for adding "_raw" to types and terms *}
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
    94
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
fun add_raw s = s ^ "_raw"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
fun add_raws ss = map add_raw ss
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
fun raw_bind bn = Binding.suffix_name "_raw" bn
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
fun replace_str ss s = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  case (AList.lookup (op=) ss s) of 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
     SOME s' => s'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
   | NONE => s
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
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
   106
  | replace_typ ty_ss T = T  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
fun raw_dts ty_ss dts =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  fun raw_dts_aux1 (bind, tys, mx) =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
    (raw_bind bind, map (replace_typ ty_ss) tys, mx)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
    (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  map raw_dts_aux2 dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  | replace_aterm trm_ss trm = trm
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
fun replace_term trm_ss ty_ss trm =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
fun rawify_dts dt_names dts dts_env =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  val raw_dts = raw_dts dts_env dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  val raw_dt_names = add_raws dt_names
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  (raw_dt_names, raw_dts)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
fun rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  val bn_funs' = map (fn (bn, ty, mx) => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
    (raw_bind bn, replace_typ dts_env ty, mx)) bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  val bn_eqs' = map (fn (attr, trm) => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
    (attr, replace_term (cnstrs_env @ bn_fun_env) dts_env trm)) bn_eqs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  (bn_funs', bn_eqs') 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
fun rawify_bclauses dts_env cnstrs_env bn_fun_env bclauses =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  fun rawify_bnds bnds = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
    map (apfst (Option.map (replace_term (cnstrs_env @ bn_fun_env) dts_env))) bnds
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  fun rawify_bclause (BEmy n) = BEmy n
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
    | rawify_bclause (BLst (bnds, bdys)) = BLst (rawify_bnds bnds, bdys)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
    | rawify_bclause (BSet (bnds, bdys)) = BSet (rawify_bnds bnds, bdys)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
    | rawify_bclause (BRes (bnds, bdys)) = BRes (rawify_bnds bnds, bdys)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  map (map (map rawify_bclause)) bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
2027
68b2d2d7b4ed Comment
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2025
diff changeset
   165
(* strip_bn_fun takes a bn function defined by the user as a union or
68b2d2d7b4ed Comment
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2025
diff changeset
   166
   append of elements and returns those elements together with bn functions
68b2d2d7b4ed Comment
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2025
diff changeset
   167
   applied *)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
fun strip_bn_fun t =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  case t of
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
    Const (@{const_name sup}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  | Const (@{const_name append}, _) $ l $ r => strip_bn_fun l @ strip_bn_fun r
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  | Const (@{const_name insert}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
      (i, NONE) :: strip_bn_fun y
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  | Const (@{const_name Cons}, _) $ (Const (@{const_name atom}, _) $ Bound i) $ y =>
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
      (i, NONE) :: strip_bn_fun y
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  | Const (@{const_name bot}, _) => []
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  | Const (@{const_name Nil}, _) => []
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
  | (f as Free _) $ Bound i => [(i, SOME f)]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
  | _ => error ("Unsupported binding function: " ^ (PolyML.makestring t))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
fun find [] _ = error ("cannot find element")
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  | 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
   186
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
ML {*
2122
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   189
fun prep_bn lthy dt_names dts eqs = 
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  fun aux eq = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
    val (lhs, rhs) = eq
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
      |> strip_qnt_body "all" 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
      |> HOLogic.dest_Trueprop
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
      |> HOLogic.dest_eq
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
    val (bn_fun, [cnstr]) = strip_comb lhs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
    val (_, ty) = dest_Free bn_fun
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
    val (ty_name, _) = dest_Type (domain_type ty)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
    val dt_index = find_index (fn x => x = ty_name) dt_names
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
    val (cnstr_head, cnstr_args) = strip_comb cnstr
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
    val rhs_elements = strip_bn_fun rhs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
    val included = map (apfst (fn i => length (cnstr_args) - i - 1)) rhs_elements
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
    (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
   206
  end
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   207
  fun aux2 eq = 
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   208
  let
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   209
    val (lhs, rhs) = eq
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   210
      |> strip_qnt_body "all" 
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   211
      |> HOLogic.dest_Trueprop
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   212
      |> HOLogic.dest_eq
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   213
    val (bn_fun, [cnstr]) = strip_comb lhs
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   214
    val (_, ty) = dest_Free bn_fun
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   215
    val (ty_name, _) = dest_Type (domain_type ty)
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   216
    val dt_index = find_index (fn x => x = ty_name) dt_names
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   217
    val (cnstr_head, cnstr_args) = strip_comb cnstr
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   218
    val rhs_elements = strip_bn_fun rhs
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   219
    val included = map (apfst (fn i => length (cnstr_args) - i - 1)) rhs_elements
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   220
  in
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   221
    (bn_fun, dt_index, (cnstr_head, included))
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  fun order dts i ts = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
    val dt = nth dts i
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
    val cts = map (fn (x, _, _) => Binding.name_of x) ((fn (_, _, _, x) => x) dt)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
    val ts' = map (fn (x, y) => (fst (dest_Const x), y)) ts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
    map (find ts') cts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  val unordered = AList.group (op=) (map aux eqs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
  val unordered' = map (fn (x, y) =>  (x, AList.group (op=) y)) unordered
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  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
   235
  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
   236
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   237
  val _ = tracing ("eqs\n" ^ cat_lines (map (Syntax.string_of_term lthy) eqs))
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   238
  val _ = tracing ("map eqs\n" ^ @{make_string} (map aux2 eqs))
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   239
  val _ = tracing ("ordered'\n" ^ @{make_string} ordered')
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
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
   241
  ordered'
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
2138
2e514a0aca63 started a new file for the parser to make some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   245
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
fun raw_nominal_decls dts bn_funs bn_eqs binds lthy =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  val thy = ProofContext.theory_of lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  val thy_name = Context.theory_name thy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  val dt_names = map (fn (_, s, _, _) => Binding.name_of s) dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  val dt_full_names = map (Long_Name.qualify thy_name) dt_names 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  val dt_full_names' = add_raws dt_full_names
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  val dts_env = dt_full_names ~~ dt_full_names'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  val cnstrs = get_cnstr_strs dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  val cnstrs_ty = get_typed_cnstrs dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  val cnstrs_full_names = map (Long_Name.qualify thy_name) cnstrs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  val cnstrs_full_names' = map (fn (x, y) => Long_Name.qualify thy_name 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
    (Long_Name.qualify (add_raw x) (add_raw y))) cnstrs_ty
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  val cnstrs_env = cnstrs_full_names ~~ cnstrs_full_names'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  val bn_fun_strs = get_bn_fun_strs bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
  val bn_fun_strs' = add_raws bn_fun_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  val bn_fun_env = bn_fun_strs ~~ bn_fun_strs'
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  val bn_fun_full_env = map (pairself (Long_Name.qualify thy_name)) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
    (bn_fun_strs ~~ bn_fun_strs')
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  val (raw_dt_names, raw_dts) = rawify_dts dt_names dts dts_env
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  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
   273
   
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  val raw_bclauses = rawify_bclauses dts_env cnstrs_env bn_fun_full_env binds 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
2122
24ca435ead14 added term4 back to the examples
Christian Urban <urbanc@in.tum.de>
parents: 2119
diff changeset
   276
  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
   277
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
  lthy 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
  |> add_datatype_wrapper raw_dt_names raw_dts 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
  ||>> add_primrec_wrapper raw_bn_funs raw_bn_eqs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  ||>> pair raw_bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  ||>> pair raw_bns
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
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
   286
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
   287
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
   288
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
   289
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
   290
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
   291
  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
   292
  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
   293
  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
   294
  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
   295
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
   296
  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
   297
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
   298
*}
6a4049e1d68d Add explicit cheats in NewParser and comment out the examples for outside use.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2016
diff changeset
   299
2023
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   300
ML {* val cheat_equivp = Unsynchronized.ref false *}
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   301
ML {* val cheat_fv_rsp = Unsynchronized.ref false *}
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   302
ML {* val cheat_alpha_bn_rsp = Unsynchronized.ref false *}
e32ec6e61154 NewParser
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2022
diff changeset
   303
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
   304
2025
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   305
ML {*
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   306
fun remove_loop t =
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   307
  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
   308
  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
   309
*}
070ba8972e97 Register only non-looping rules in eq_iff
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2023
diff changeset
   310
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
   311
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
   312
  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
   313
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
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
   315
  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
   316
  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
   317
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
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
   319
  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
   320
     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
   321
      
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
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
   323
  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
   324
  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
   325
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
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
   327
  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
   328
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
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
   330
  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
   331
     (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
   332
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
   333
  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
   334
  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
   335
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
   336
  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
   337
     (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
   338
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
   339
 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
   340
     (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
   341
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
   342
 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
   343
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
   344
 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
   345
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
   346
 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
   347
 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
   348
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
   349
 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
   350
     (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
   351
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
   352
 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
   353
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
   354
 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
   355
     (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
   356
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
   357
 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
   358
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
   359
 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
   360
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
   361
 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
   362
 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
   363
 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
   364
 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
   365
 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
   366
 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
   367
 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
   368
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
   369
 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
   370
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
   371
 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
   372
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
   373
 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
   374
*}
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
   375
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
   376
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
   377
(* 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
   378
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
   379
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
   380
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
   381
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
   382
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
   383
*}
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
   384
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
   385
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
   386
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
   387
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
fun nominal_datatype2 dts bn_funs bn_eqs bclauses lthy =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
let
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
   391
  (* definition of the raw datatype and raw bn-functions *)
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
   392
  val ((((raw_dt_names, (raw_bn_funs, raw_bn_eqs)), raw_bclauses), raw_bns), 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
   393
    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
   394
    else raise TEST lthy
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
1945
93e5a31ba230 added definition of raw-permutations to the new-parser
Christian Urban <urbanc@in.tum.de>
parents: 1944
diff changeset
   396
  val dtinfo = Datatype.the_info (ProofContext.theory_of lthy1) (hd raw_dt_names);
93e5a31ba230 added definition of raw-permutations to the new-parser
Christian Urban <urbanc@in.tum.de>
parents: 1944
diff changeset
   397
  val {descr, sorts, ...} = dtinfo;
2011
12ce87b55f97 tried to add some comments in the huge(!) nominal2_cmd function
Christian Urban <urbanc@in.tum.de>
parents: 2008
diff changeset
   398
  val raw_tys = map (fn (i, _) => nth_dtyp descr sorts i) descr;
1945
93e5a31ba230 added definition of raw-permutations to the new-parser
Christian Urban <urbanc@in.tum.de>
parents: 1944
diff changeset
   399
2047
31ba33a199c7 fixed my error with define_raw_fv
Christian Urban <urbanc@in.tum.de>
parents: 2046
diff changeset
   400
  val induct_thm = #induct dtinfo;
1999
4df3441aa0b4 more parser/new parser synchronization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1996
diff changeset
   401
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
   402
  (* definitions of raw permutations *)
1945
93e5a31ba230 added definition of raw-permutations to the new-parser
Christian Urban <urbanc@in.tum.de>
parents: 1944
diff changeset
   403
  val ((raw_perm_def, raw_perm_simps, perms), lthy2) =
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
   404
    if get_STEPS lthy > 2 
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
   405
    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
   406
    else raise TEST lthy1
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1945
diff changeset
   407
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
   408
  (* 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
   409
  val morphism_2_0 = ProofContext.export_morphism lthy2 lthy
2118
0e52851acac4 moved the data-transformation into the parser
Christian Urban <urbanc@in.tum.de>
parents: 2107
diff changeset
   410
  fun export_fun f (t, n , l) = (f t, n, map (map (apsnd (Option.map f))) l);
0e52851acac4 moved the data-transformation into the parser
Christian Urban <urbanc@in.tum.de>
parents: 2107
diff changeset
   411
  val bn_funs_decls = map (export_fun (Morphism.term morphism_2_0)) raw_bns;
0e52851acac4 moved the data-transformation into the parser
Christian Urban <urbanc@in.tum.de>
parents: 2107
diff changeset
   412
 
2138
2e514a0aca63 started a new file for the parser to make some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   413
  val thy = Local_Theory.exit_global lthy2;
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
   414
  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
   415
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1945
diff changeset
   416
  val lthy3 = Theory_Target.init NONE thy;
2000
f18b8e8a4909 Merged nominal_datatype into NewParser until eqvts
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1999
diff changeset
   417
  val raw_bn_funs = map (fn (f, _, _) => f) bn_funs_decls;
1960
47e2e91705f3 Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1945
diff changeset
   418
2106
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
   419
  val _ = tracing ("raw_bns\n" ^ @{make_string} raw_bns)
409ecb7284dd properly exported defined bn-functions
Christian Urban <urbanc@in.tum.de>
parents: 2105
diff changeset
   420
  val _ = tracing ("bn_funs\n" ^ @{make_string} bn_funs_decls)
2118
0e52851acac4 moved the data-transformation into the parser
Christian Urban <urbanc@in.tum.de>
parents: 2107
diff changeset
   421
   
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
   422
  val ((fv, fvbn), fv_def, lthy3a) = 
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
   423
    if get_STEPS lthy > 3 
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
   424
    then define_raw_fv descr sorts bn_funs_decls 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
   425
    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
   426
  
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
in
2138
2e514a0aca63 started a new file for the parser to make some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2125
diff changeset
   428
  (0, 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
   429
end handle TEST ctxt => (0, ctxt)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
section {* Preparing and parsing of the specification *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
ML {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
(* parsing the datatypes and declaring *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
(* constructors in the local theory    *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
fun prepare_dts dt_strs lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  val thy = ProofContext.theory_of lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
  fun mk_type full_tname tvrs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
    Type (full_tname, map (fn a => TVar ((a, 0), [])) tvrs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
  fun prep_cnstr full_tname tvs (cname, anno_tys, mx, _) =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
    val tys = map (Syntax.read_typ lthy o snd) anno_tys
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
    val ty = mk_type full_tname tvs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
    ((cname, tys ---> ty, mx), (cname, tys, mx))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
  fun prep_dt (tvs, tname, mx, cnstrs) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
    val full_tname = Sign.full_name thy tname
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
    val (cnstrs', cnstrs'') = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
      split_list (map (prep_cnstr full_tname tvs) cnstrs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
    (cnstrs', (tvs, tname, mx, cnstrs''))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
  end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
  val (cnstrs, dts) = split_list (map prep_dt dt_strs)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
  lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
  |> pair dts
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
(* parsing the binding function specification and *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
(* declaring the functions in the local theory    *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
  val ((bn_funs, bn_eqs), _) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
    Specification.read_spec bn_fun_strs bn_eq_strs lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
  fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
  val bn_funs' = map prep_bn_fun bn_funs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
  lthy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
  |> Local_Theory.theory (Sign.add_consts_i bn_funs')
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
  |> pair (bn_funs', bn_eqs) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
end 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
text {* associates every SOME with the index in the list; drops NONEs *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
fun indexify xs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
  fun mapp _ [] = []
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
    | mapp i (NONE :: xs) = mapp (i + 1) xs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
    | mapp i (SOME x :: xs) = (x, i) :: mapp (i + 1) xs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
in 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
  mapp 0 xs 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
fun index_lookup xs x =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
  case AList.lookup (op=) xs x of
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
    SOME x => x
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
  | NONE => error ("Cannot find " ^ x ^ " as argument annotation.");
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
fun prepare_bclauses dt_strs lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
  val annos_bclauses =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
    get_cnstrs dt_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
    |> map (map (fn (_, antys, _, bns) => (map fst antys, bns)))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
  fun prep_binder env bn_str =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
    case (Syntax.read_term lthy bn_str) of
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
      Free (x, _) => (NONE, index_lookup env x)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
    | Const (a, T) $ Free (x, _) => (SOME (Const (a, T)), index_lookup env x)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
    | _ => error ("The term " ^ bn_str ^ " is not allowed as binding function.")
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
  fun prep_body env bn_str = index_lookup env bn_str
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
  fun prep_mode "bind"     = BLst 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
    | prep_mode "bind_set" = BSet 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
    | prep_mode "bind_res" = BRes 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  fun prep_bclause env (mode, binders, bodies) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
    val binders' = map (prep_binder env) binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
    val bodies' = map (prep_body env) bodies
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
  in  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
    prep_mode mode (binders', bodies')
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
  fun prep_bclauses (annos, bclause_strs) = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
    val env = indexify annos (* for every label, associate the index *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
    map (prep_bclause env) bclause_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
  map (map prep_bclauses) annos_bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
1943
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   542
text {* 
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   543
  adds an empty binding clause for every argument
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   544
  that is not already part of a binding clause
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   545
*}
48add8d4d7e5 slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 1942
diff changeset
   546
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
fun included i bcs = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  fun incl (BEmy j) = i = j
2076
6cb1a4639521 mem => member
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2048
diff changeset
   551
    | incl (BLst (bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
6cb1a4639521 mem => member
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2048
diff changeset
   552
    | incl (BSet (bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
6cb1a4639521 mem => member
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2048
diff changeset
   553
    | incl (BRes (bns, bds)) = (member (op =) (map snd bns) i) orelse (member (op =) bds i)
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
  exists incl bcs 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
*}
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
fun complete dt_strs bclauses = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
  val args = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
    get_cnstrs dt_strs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
    |> map (map (fn (_, antys, _, _) => length antys))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
  fun complt n bcs = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
  let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
    fun add bcs i = (if included i bcs then [] else [BEmy i]) 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
  in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
    bcs @ (flat (map_range (add bcs) n))
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
  end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
  map2 (map2 complt) args bclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
ML {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
  val lthy0 = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
    Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) lthy
1944
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
   583
  val (dts, lthy1) = prepare_dts dt_strs lthy0
Christian Urban <urbanc@in.tum.de>
parents: 1943
diff changeset
   584
  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
   585
  val bclauses = prepare_bclauses dt_strs lthy2
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
  val bclauses' = complete dt_strs bclauses 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
  nominal_datatype2 dts bn_funs bn_eqs bclauses' lthy |> snd
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
(* Command Keyword *)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
val _ = OuterSyntax.local_theory "nominal_datatype" "test" OuterKeyword.thy_decl
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
  (main_parser >> nominal_datatype2_cmd)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
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
   598
(*
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
   599
atom_decl name
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
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
   601
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
   602
  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
   603
| 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
   604
| 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
   605
| 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
   606
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
   607
  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
   608
| 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
   609
binder
1981
9f9c4965b608 Include support of unknown datatypes in new fv
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1971
diff changeset
   610
 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
   611
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
   612
  "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
   613
| "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
   614
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
   615
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
   616
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
   617
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
   618
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
   619
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
   620
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
   621
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
   622
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
   623
thm lam_pt.perm
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
nominal_datatype exp =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
  EVar name
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
| EUnit
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   628
| EPair q1::exp q2::exp
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
| ELetRec l::lrbs e::exp     bind "b_lrbs l" in e l
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
and fnclause =
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   632
  K x::name p::pat f::exp    bind_res "b_pat p" in f
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
and fnclauses =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
  S fnclause
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
| ORs fnclause fnclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
and lrb =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
  Clause fnclauses
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
and lrbs =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
  Single lrb
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
| More lrb lrbs
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
and pat =
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   646
  PVar name
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
| PUnit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
| PPair pat pat
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
binder
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   651
  b_lrbs      :: "lrbs \<Rightarrow> atom list" and
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
  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
   653
  b_fnclauses :: "fnclauses \<Rightarrow> atom list" and
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   654
  b_fnclause  :: "fnclause \<Rightarrow> atom list" and
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   655
  b_lrb       :: "lrb \<Rightarrow> atom list"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
where
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
  "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
   659
| "b_lrbs (More l ls) = append (b_lrb l) (b_lrbs ls)"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
| "b_pat (PVar x) = {atom x}"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
| "b_pat (PUnit) = {}"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
| "b_pat (PPair p1 p2) = b_pat p1 \<union> b_pat p2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
| "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
   664
| "b_fnclauses (ORs fc fcs) = append (b_fnclause fc) (b_fnclauses fcs)"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
| "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
   666
| "b_fnclause (K x pat exp) = [atom x]"
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
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
   668
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
   669
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
   670
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
   671
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
   672
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
   673
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   674
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
   675
  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
   676
| 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
   677
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
   678
  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
   679
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
   680
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
   681
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
   682
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
   683
*)
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
   684
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
   685
(* 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
   686
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
   687
(*
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
   688
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
   689
  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
   690
| Fn2 "ty2" "ty2"
1988
4444d52201dc Fixing the definitions in the Parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1981
diff changeset
   691
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
   692
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
   693
  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
   694
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
   695
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
   696
  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
   697
| 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
   698
| 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
   699
*)
1964
209ee65b2395 added some further problemetic tests
Christian Urban <urbanc@in.tum.de>
parents: 1963
diff changeset
   700
1941
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
end
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706