src/HOL/Datatype_Examples/Lift_BNF.thy
author nipkow
Mon, 15 Jan 2024 22:50:13 +0100
changeset 79490 b287510a4202
parent 71469 d7ef73df3d15
child 79564 33b10cd883ae
permissions -rw-r--r--
Added time function automation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62695
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     1
(*  Title:      HOL/Datatype_Examples/Lift_BNF.thy
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     2
    Author:     Dmitriy Traytel, ETH Zürich
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     3
    Copyright   2015
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     4
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     5
Demonstration of the "lift_bnf" command.
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     6
*)
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     7
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62697
diff changeset
     8
section \<open>Demonstration of the \textbf{lift_bnf} Command\<close>
62695
b287b56a6ce5 file header
blanchet
parents: 62137
diff changeset
     9
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    10
theory Lift_BNF
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    11
imports Main
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    12
begin
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    13
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    14
subsection \<open>Lifting \textbf{typedef}s\<close>
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    15
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    16
typedef 'a nonempty_list = "{xs :: 'a list. xs \<noteq> []}"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    17
  by blast
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    18
setup_lifting type_definition_nonempty_list
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    19
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    20
lift_bnf (no_warn_wits) (neset: 'a) nonempty_list
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    21
  for map: nemap rel: nerel
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    22
  by auto
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    23
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    24
typedef ('a :: finite, 'b) fin_nonempty_list = "{(xs :: 'a set, ys :: 'b list). ys \<noteq> []}"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    25
  by blast
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    26
setup_lifting type_definition_fin_nonempty_list
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    27
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    28
lift_bnf (no_warn_wits) (dead 'a :: finite, 'b) fin_nonempty_list
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    29
  [wits: "\<lambda>b. ({} :: 'a :: finite set, [b :: 'b])"]
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    30
  by auto
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    31
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    32
datatype 'a tree = Leaf | Node 'a "'a tree nonempty_list"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    33
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    34
record 'a point =
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    35
  xCoord :: 'a
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    36
  yCoord :: 'a
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    37
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    38
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    39
copy_bnf (no_warn_transfer) ('a, 's) point_ext
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    40
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    41
typedef 'a it = "UNIV :: 'a set"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    42
  by blast
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    43
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    44
setup_lifting type_definition_it
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    45
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    46
copy_bnf (plugins del: size) 'a it
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    47
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    48
typedef ('a, 'b) T_prod = "UNIV :: ('a \<times> 'b) set"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    49
  by blast
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    50
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    51
setup_lifting type_definition_T_prod
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    52
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    53
copy_bnf ('a, 'b) T_prod
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    54
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    55
typedef ('a, 'b, 'c) T_func = "UNIV :: ('a \<Rightarrow> 'b * 'c) set"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    56
  by blast
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    57
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    58
setup_lifting type_definition_T_func
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    59
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    60
copy_bnf ('a, 'b, 'c) T_func
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    61
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    62
typedef ('a, 'b) sum_copy = "UNIV :: ('a + 'b) set"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    63
  by blast
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    64
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    65
setup_lifting type_definition_sum_copy
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    66
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    67
copy_bnf ('a, 'b) sum_copy
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    68
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    69
typedef ('a, 'b) T_sum = "{Inl x | x. True} :: ('a + 'b) set"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    70
  by blast
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    71
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    72
lift_bnf (no_warn_wits, no_warn_transfer) ('a, 'b) T_sum [wits: "Inl :: 'a \<Rightarrow> 'a + 'b"]
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    73
  by (force simp: map_sum_def sum_set_defs split: sum.splits)+
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    74
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    75
typedef ('key, 'value) alist = "{xs :: ('key \<times> 'value) list. (distinct \<circ> map fst) xs}"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    76
  morphisms impl_of Alist
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    77
proof
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    78
  show "[] \<in> {xs. (distinct o map fst) xs}"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    79
    by simp
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    80
qed
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    81
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    82
setup_lifting type_definition_alist
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    83
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    84
lift_bnf (dead 'k, 'v) alist [wits: "Nil :: ('k \<times> 'v) list"]
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    85
  by auto
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    86
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    87
typedef 'a myopt = "{X :: 'a set. finite X \<and> card X \<le> 1}" by (rule exI[of _ "{}"]) auto
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    88
lemma myopt_type_def: "type_definition
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    89
  (\<lambda>X. if card (Rep_myopt X) = 1 then Some (the_elem (Rep_myopt X)) else None)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    90
  (\<lambda>x. Abs_myopt (case x of Some x \<Rightarrow> {x} | _ \<Rightarrow> {}))
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    91
  (UNIV :: 'a option set)"
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    92
  apply unfold_locales
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    93
    apply (auto simp: Abs_myopt_inverse dest!: card_eq_SucD split: option.splits)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    94
   apply (metis Rep_myopt_inverse)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    95
  apply (metis One_nat_def Rep_myopt Rep_myopt_inverse Suc_le_mono card_0_eq le0 le_antisym mem_Collect_eq nat.exhaust)
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    96
  done
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    97
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
    98
copy_bnf (no_warn_transfer) 'a myopt via myopt_type_def
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
    99
62137
b8dc1fd7d900 more careful witness' type analysis
traytel
parents: 60918
diff changeset
   100
typedef ('k, 'v) fmap = "{M :: ('k \<rightharpoonup> 'v). finite (dom M)}"
b8dc1fd7d900 more careful witness' type analysis
traytel
parents: 60918
diff changeset
   101
  by (rule exI[of _ Map.empty]) simp_all
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   102
setup_lifting type_definition_fmap
62137
b8dc1fd7d900 more careful witness' type analysis
traytel
parents: 60918
diff changeset
   103
b8dc1fd7d900 more careful witness' type analysis
traytel
parents: 60918
diff changeset
   104
lift_bnf (dead 'k, 'v) fmap [wits: "Map.empty :: 'k \<Rightarrow> 'v option"]
b8dc1fd7d900 more careful witness' type analysis
traytel
parents: 60918
diff changeset
   105
  by auto
b8dc1fd7d900 more careful witness' type analysis
traytel
parents: 60918
diff changeset
   106
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   107
typedef ('a, 'b) tuple_dead = "UNIV :: ('a \<times> 'b) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   108
typedef ('a, 'b) tuple_dead1 = "UNIV :: ('a \<times> 'b) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   109
typedef ('a, 'b) tuple_dead2 = "UNIV :: ('a \<times> 'b) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   110
typedef ('a, 'b, 'c) triple_dead = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   111
typedef ('a, 'b, 'c) triple_dead1 = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   112
typedef ('a, 'b, 'c) triple_dead2 = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   113
typedef ('a, 'b, 'c) triple_dead3 = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   114
typedef ('a, 'b, 'c) triple_dead12 = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   115
typedef ('a, 'b, 'c) triple_dead13 = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   116
typedef ('a, 'b, 'c) triple_dead23 = "UNIV :: ('a \<times> 'b \<times> 'c) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   117
typedef ('a, 'b, 'c, 'd) quadruple_dead =    "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   118
typedef ('a, 'b, 'c, 'd) quadruple_dead1 =   "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   119
typedef ('a, 'b, 'c, 'd) quadruple_dead2 =   "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   120
typedef ('a, 'b, 'c, 'd) quadruple_dead3 =   "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   121
typedef ('a, 'b, 'c, 'd) quadruple_dead4 =   "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   122
typedef ('a, 'b, 'c, 'd) quadruple_dead12 =  "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   123
typedef ('a, 'b, 'c, 'd) quadruple_dead13 =  "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   124
typedef ('a, 'b, 'c, 'd) quadruple_dead14 =  "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   125
typedef ('a, 'b, 'c, 'd) quadruple_dead23 =  "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   126
typedef ('a, 'b, 'c, 'd) quadruple_dead24 =  "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   127
typedef ('a, 'b, 'c, 'd) quadruple_dead34 =  "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   128
typedef ('a, 'b, 'c, 'd) quadruple_dead123 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   129
typedef ('a, 'b, 'c, 'd) quadruple_dead124 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   130
typedef ('a, 'b, 'c, 'd) quadruple_dead134 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   131
typedef ('a, 'b, 'c, 'd) quadruple_dead234 = "UNIV :: ('a \<times> 'b \<times> 'c \<times> 'd) set" ..
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   132
setup_lifting type_definition_tuple_dead
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   133
setup_lifting type_definition_tuple_dead1
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   134
setup_lifting type_definition_tuple_dead2
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   135
setup_lifting type_definition_triple_dead
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   136
setup_lifting type_definition_triple_dead1
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   137
setup_lifting type_definition_triple_dead2
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   138
setup_lifting type_definition_triple_dead3
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   139
setup_lifting type_definition_triple_dead12
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   140
setup_lifting type_definition_triple_dead13
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   141
setup_lifting type_definition_triple_dead23
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   142
setup_lifting type_definition_quadruple_dead
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   143
setup_lifting type_definition_quadruple_dead1
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   144
setup_lifting type_definition_quadruple_dead2
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   145
setup_lifting type_definition_quadruple_dead3
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   146
setup_lifting type_definition_quadruple_dead4
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   147
setup_lifting type_definition_quadruple_dead12
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   148
setup_lifting type_definition_quadruple_dead13
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   149
setup_lifting type_definition_quadruple_dead14
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   150
setup_lifting type_definition_quadruple_dead23
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   151
setup_lifting type_definition_quadruple_dead24
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   152
setup_lifting type_definition_quadruple_dead34
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   153
setup_lifting type_definition_quadruple_dead123
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   154
setup_lifting type_definition_quadruple_dead124
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   155
setup_lifting type_definition_quadruple_dead134
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   156
setup_lifting type_definition_quadruple_dead234
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   157
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   158
lift_bnf (no_warn_wits) (     'a,      'b) tuple_dead by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   159
lift_bnf (no_warn_wits) (dead 'a,      'b) tuple_dead1 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   160
lift_bnf (no_warn_wits) (     'a, dead 'b) tuple_dead2 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   161
lift_bnf (no_warn_wits) (     'a,      'b,      'c) triple_dead by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   162
lift_bnf (no_warn_wits) (dead 'a,      'b,      'c) triple_dead1 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   163
lift_bnf (no_warn_wits) (     'a, dead 'b,      'c) triple_dead2 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   164
lift_bnf (no_warn_wits) (     'a,      'b, dead 'c) triple_dead3 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   165
lift_bnf (no_warn_wits) (dead 'a, dead 'b,      'c) triple_dead12 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   166
lift_bnf (no_warn_wits) (dead 'a,      'b, dead 'c) triple_dead13 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   167
lift_bnf (no_warn_wits) (     'a, dead 'b, dead 'c) triple_dead23 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   168
lift_bnf (no_warn_wits) (     'a,      'b,      'c,      'd) quadruple_dead  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   169
lift_bnf (no_warn_wits) (dead 'a,      'b,      'c,      'd) quadruple_dead1 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   170
lift_bnf (no_warn_wits) (     'a, dead 'b,      'c,      'd) quadruple_dead2 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   171
lift_bnf (no_warn_wits) (     'a,      'b, dead 'c,      'd) quadruple_dead3 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   172
lift_bnf (no_warn_wits) (     'a,      'b,      'c, dead 'd) quadruple_dead4 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   173
lift_bnf (no_warn_wits) (dead 'a, dead 'b,      'c,      'd) quadruple_dead12  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   174
lift_bnf (no_warn_wits) (dead 'a,      'b, dead 'c,      'd) quadruple_dead13  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   175
lift_bnf (no_warn_wits) (dead 'a,      'b,      'c, dead 'd) quadruple_dead14  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   176
lift_bnf (no_warn_wits) (     'a, dead 'b, dead 'c,      'd) quadruple_dead23  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   177
lift_bnf (no_warn_wits) (     'a, dead 'b,      'c, dead 'd) quadruple_dead24  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   178
lift_bnf (no_warn_wits) (     'a,      'b, dead 'c, dead 'd) quadruple_dead34  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   179
lift_bnf (no_warn_wits) (dead 'a, dead 'b, dead 'c,      'd) quadruple_dead123 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   180
lift_bnf (no_warn_wits) (dead 'a, dead 'b,      'c, dead 'd) quadruple_dead124 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   181
lift_bnf (no_warn_wits) (dead 'a,      'b, dead 'c, dead 'd) quadruple_dead134 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   182
lift_bnf (no_warn_wits) (     'a, dead 'b, dead 'c, dead 'd) quadruple_dead234 by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   183
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   184
subsection \<open>Lifting \textbf{quotient_type}s\<close>
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   185
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   186
quotient_type 'a cpy_list = "'a list" / "(=)"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   187
  by (rule identity_equivp)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   188
71469
d7ef73df3d15 lift BNF witnesses for quotients (unless better ones are specified by the user)
traytel
parents: 71262
diff changeset
   189
lift_bnf 'a cpy_list
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   190
  by (auto intro: list_all2_trans)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   191
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   192
quotient_type ('a, 'b) funk = "('a \<Rightarrow> 'b)" / "\<lambda>f g. \<forall>a. f a = g a"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   193
  unfolding equivp_def by metis
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   194
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   195
lemma funk_closure: "{(x, x'). \<forall>a. x a = x' a} `` {x. range x \<subseteq> A} = {x. range x \<subseteq> A}"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   196
  by auto
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   197
71469
d7ef73df3d15 lift BNF witnesses for quotients (unless better ones are specified by the user)
traytel
parents: 71262
diff changeset
   198
lift_bnf (dead 'a, 'b) funk
71262
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   199
  unfolding funk_closure rel_fun_def by (auto 0 10 split: if_splits)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   200
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   201
lemma assumes "equivp REL"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   202
  shows "REL OO REL OO R = REL OO R"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   203
  using equivp_transp[OF assms] equivp_reflp[OF assms]
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   204
  by blast
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   205
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   206
inductive ignore_Inl :: "'a + 'a \<Rightarrow> 'a + 'a \<Rightarrow> bool" where
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   207
  "ignore_Inl (Inl x) (Inl y)"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   208
| "ignore_Inl (Inr x) (Inr x)"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   209
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   210
inductive_simps [simp]:
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   211
  "ignore_Inl (Inl x) y"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   212
  "ignore_Inl (Inr x') y"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   213
  "ignore_Inl y (Inl x)"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   214
  "ignore_Inl y (Inr x')"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   215
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   216
quotient_type 'a opt = "'a + 'a" / ignore_Inl
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   217
  apply(auto intro!: equivpI simp add: reflp_def symp_def transp_def elim!: ignore_Inl.cases)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   218
  subgoal for x by(cases x) simp_all.
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   219
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   220
lemma ignore_Inl_map_respect:
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   221
  "(rel_fun (=) (rel_fun ignore_Inl ignore_Inl)) (\<lambda>f. map_sum f f) (\<lambda>f. map_sum f f)"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   222
  by(auto simp add: rel_fun_def elim: ignore_Inl.cases)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   223
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   224
lemma ignore_Inl_pos_distr:
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   225
  "A OO B \<noteq> bot \<Longrightarrow> ignore_Inl OO rel_sum A A OO ignore_Inl OO rel_sum B B OO ignore_Inl \<le>
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   226
   ignore_Inl OO rel_sum (A OO B) (A OO B) OO ignore_Inl"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   227
  by(fastforce elim!: ignore_Inl.cases simp add: relcompp_apply fun_eq_iff intro: exI[where x="Inl _"])
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   228
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   229
lemma ignore_Inl_Inter:
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   230
  "\<lbrakk> S \<noteq> {}; \<Inter>S \<noteq> {} \<rbrakk> \<Longrightarrow> (\<Inter>A\<in>S. {(x, y). ignore_Inl x y} `` {x. Basic_BNFs.setl x \<union> Basic_BNFs.setr x \<subseteq> A}) \<subseteq> {(x, y). ignore_Inl x y} `` (\<Inter>A\<in>S. {x. Basic_BNFs.setl x \<union> Basic_BNFs.setr x \<subseteq> A})"
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   231
  apply(clarsimp; safe; clarsimp)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   232
  subgoal for x A x'
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   233
    apply(cases x)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   234
     apply(drule bspec[where x="Inl x'"])
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   235
      apply clarsimp
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   236
     apply simp
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   237
    apply clarsimp
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   238
    apply(drule bspec[where x="Inr _"])
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   239
     apply simp
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   240
    apply simp
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   241
    done
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   242
  done
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   243
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   244
lift_bnf 'a opt [wits: "(Inl undefined :: 'a + 'a)"]
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   245
   apply(auto simp add: rel_fun_def elim: ignore_Inl.cases)[]
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   246
  apply(fastforce simp add: rel_sum.simps)
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   247
  subgoal for Ss
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   248
    using ignore_Inl_Inter[unfolded Plus_def, of Ss]
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   249
    by blast
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   250
  apply (auto simp: Ball_def image_iff sum_set_defs elim: ignore_Inl.cases split: sum.splits) []
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   251
  done
a30278c8585f extension of lift_bnf to support quotient types
traytel
parents: 63167
diff changeset
   252
60918
4ceef1592e8c new command for lifting BNF structure over typedefs
traytel
parents:
diff changeset
   253
end