src/HOL/Analysis/measurable.ML
author wenzelm
Sat, 08 May 2021 13:06:30 +0200
changeset 73652 d5c3eee7da74
parent 73551 53c148e39819
child 74152 069f6b2c5a07
permissions -rw-r--r--
separate component for idea-icons.jar, from jedit_build (see also ff0e0bb81597);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     1
(*  Title:      HOL/Analysis/measurable.ML
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl <hoelzl@in.tum.de>
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     3
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     4
Measurability prover.
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     5
*)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     6
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     7
signature MEASURABLE = 
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     8
sig
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
     9
  type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    10
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    11
  datatype level = Concrete | Generic
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    12
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    13
  val dest_thm_attr : attribute context_parser
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    14
  val cong_thm_attr : attribute context_parser
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    15
  val measurable_thm_attr : bool * (bool * level) -> attribute
53043
8cbfbeb566a4 more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
wenzelm
parents: 51717
diff changeset
    16
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    17
  val add_del_cong_thm : bool -> thm -> Context.generic -> Context.generic ;
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    18
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    19
  val get_all : Context.generic -> thm list
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    20
  val get_dest : Context.generic -> thm list
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    21
  val get_cong : Context.generic -> thm list
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    22
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    23
  val measurable_tac : Proof.context -> thm list -> tactic
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    24
53043
8cbfbeb566a4 more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
wenzelm
parents: 51717
diff changeset
    25
  val simproc : Proof.context -> cterm -> thm option
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    26
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    27
  val add_preprocessor : string -> preprocessor -> Context.generic -> Context.generic
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    28
  val del_preprocessor : string -> Context.generic -> Context.generic
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    29
  val add_local_cong : thm -> Proof.context -> Proof.context
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    30
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    31
  val prepare_facts : Proof.context -> thm list -> (thm list * Proof.context)
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    32
end ;
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    33
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    34
structure Measurable : MEASURABLE =
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    35
struct
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    36
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    37
type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    38
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    39
datatype level = Concrete | Generic;
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    40
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    41
fun eq_measurable_thms ((th1, d1), (th2, d2)) = 
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    42
  d1 = d2 andalso Thm.eq_thm_prop (th1, th2) ;
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    43
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    44
fun merge_dups (xs:(string * preprocessor) list) ys =
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    45
  xs @ (filter (fn (name, _) => is_none (find_first (fn (name', _) => name' = name) xs)) ys) 
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    46
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    47
structure Data = Generic_Data
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    48
(
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    49
  type T = {
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    50
    measurable_thms : (thm * (bool * level)) Item_Net.T,
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    51
    dest_thms : thm Item_Net.T,
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    52
    cong_thms : thm Item_Net.T,
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    53
    preprocessors : (string * preprocessor) list }
59992
d8db5172c23f make SML/NJ more happy;
wenzelm
parents: 59837
diff changeset
    54
  val empty: T = {
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    55
    measurable_thms = Item_Net.init eq_measurable_thms (single o Thm.prop_of o fst),
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    56
    dest_thms = Thm.full_rules,
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    57
    cong_thms = Thm.full_rules,
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    58
    preprocessors = [] };
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    59
  val extend = I;
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    60
  fun merge ({measurable_thms = t1, dest_thms = dt1, cong_thms = ct1, preprocessors = i1 },
59992
d8db5172c23f make SML/NJ more happy;
wenzelm
parents: 59837
diff changeset
    61
      {measurable_thms = t2, dest_thms = dt2, cong_thms = ct2, preprocessors = i2 }) : T = {
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    62
    measurable_thms = Item_Net.merge (t1, t2),
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    63
    dest_thms = Item_Net.merge (dt1, dt2),
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    64
    cong_thms = Item_Net.merge (ct1, ct2),
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    65
    preprocessors = merge_dups i1 i2 
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    66
    };
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    67
);
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    68
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    69
val debug =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
    70
  Attrib.setup_config_bool \<^binding>\<open>measurable_debug\<close> (K false)
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    71
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    72
val split =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
    73
  Attrib.setup_config_bool \<^binding>\<open>measurable_split\<close> (K true)
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    74
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    75
fun map_data f1 f2 f3 f4
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    76
  {measurable_thms = t1,    dest_thms = t2,    cong_thms = t3,    preprocessors = t4 } =
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    77
  {measurable_thms = f1 t1, dest_thms = f2 t2, cong_thms = f3 t3, preprocessors = f4 t4}
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    78
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    79
fun map_measurable_thms f = map_data f I I I
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    80
fun map_dest_thms f = map_data I f I I
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    81
fun map_cong_thms f = map_data I I f I
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
    82
fun map_preprocessors f = map_data I I I f
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    83
59992
d8db5172c23f make SML/NJ more happy;
wenzelm
parents: 59837
diff changeset
    84
fun generic_add_del map : attribute context_parser =
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    85
  Scan.lift
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    86
    (Args.add >> K Item_Net.update || Args.del >> K Item_Net.remove || Scan.succeed Item_Net.update) >>
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    87
    (fn f => Thm.declaration_attribute (Data.map o map o f))
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    88
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    89
val dest_thm_attr = generic_add_del map_dest_thms
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    90
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    91
val cong_thm_attr = generic_add_del map_cong_thms
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    92
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    93
fun del_thm th net =
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    94
  let
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    95
    val thms = net |> Item_Net.content |> filter (fn (th', _) => Thm.eq_thm (th, th'))
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    96
  in fold Item_Net.remove thms net end ;
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    97
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    98
fun measurable_thm_attr (do_add, d) = Thm.declaration_attribute
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
    99
  (Data.map o map_measurable_thms o (if do_add then Item_Net.update o rpair d else del_thm))
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   100
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   101
val get_dest = Item_Net.content o #dest_thms o Data.get;
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   102
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   103
val get_cong = Item_Net.content o #cong_thms o Data.get;
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   104
val add_cong = Data.map o map_cong_thms o Item_Net.update;
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   105
val del_cong = Data.map o map_cong_thms o Item_Net.remove;
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   106
fun add_del_cong_thm true = add_cong
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   107
  | add_del_cong_thm false = del_cong
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   108
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   109
fun add_preprocessor name f = Data.map (map_preprocessors (fn xs => xs @ [(name, f)]))
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   110
fun del_preprocessor name = Data.map (map_preprocessors (filter (fn (n, _) => n <> name)))
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   111
val add_local_cong = Context.proof_map o add_cong
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   112
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   113
val get_preprocessors = Context.Proof #> Data.get #> #preprocessors ;
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   114
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   115
fun is_too_generic thm =
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   116
  let 
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   117
    val concl = Thm.concl_of thm
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   118
    val concl' = HOLogic.dest_Trueprop concl handle TERM _ => concl
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   119
  in is_Var (head_of concl') end
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   120
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   121
val get_thms = Data.get #> #measurable_thms #> Item_Net.content ;
58965
a62cdcc5344b add del option to measurable;
Andreas Lochbihler
parents: 56491
diff changeset
   122
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   123
val get_all = get_thms #> map fst ;
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   124
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   125
fun debug_tac ctxt msg f =
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   126
  if Config.get ctxt debug then print_tac ctxt (msg ()) THEN f else f
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   127
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   128
fun nth_hol_goal thm i =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   129
  HOLogic.dest_Trueprop (Logic.strip_imp_concl (strip_all_body (nth (Thm.prems_of thm) (i - 1))))
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   130
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   131
fun dest_measurable_fun t =
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   132
  (case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   133
    (Const (\<^const_name>\<open>Set.member\<close>, _) $ f $ (Const (\<^const_name>\<open>measurable\<close>, _) $ _ $ _)) => f
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   134
  | _ => raise (TERM ("not a measurability predicate", [t])))
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   135
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   136
fun not_measurable_prop n thm =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   137
  if length (Thm.prems_of thm) < n then false
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   138
  else
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   139
    (case nth_hol_goal thm n of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   140
      (Const (\<^const_name>\<open>Set.member\<close>, _) $ _ $ (Const (\<^const_name>\<open>sets\<close>, _) $ _)) => false
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   141
    | (Const (\<^const_name>\<open>Set.member\<close>, _) $ _ $ (Const (\<^const_name>\<open>measurable\<close>, _) $ _ $ _)) => false
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   142
    | _ => true)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   143
    handle TERM _ => true;
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   144
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   145
fun indep (Bound i) t b = i < b orelse t <= i
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   146
  | indep (f $ t) top bot = indep f top bot andalso indep t top bot
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   147
  | indep (Abs (_,_,t)) top bot = indep t (top + 1) (bot + 1)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   148
  | indep _ _ _ = true;
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   149
59837
wenzelm
parents: 59621
diff changeset
   150
fun cnt_prefixes ctxt (Abs (n, T, t)) =
wenzelm
parents: 59621
diff changeset
   151
    let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   152
      fun is_countable ty = Sign.of_sort (Proof_Context.theory_of ctxt) (ty, \<^sort>\<open>countable\<close>)
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   153
      fun cnt_walk (Abs (ns, T, t)) Ts =
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   154
          map (fn (t', t'') => (Abs (ns, T, t'), t'')) (cnt_walk t (T::Ts))
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   155
        | cnt_walk (f $ g) Ts = let
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   156
            val n = length Ts - 1
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   157
          in
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   158
            map (fn (f', t) => (f' $ g, t)) (cnt_walk f Ts) @
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   159
            map (fn (g', t) => (f $ g', t)) (cnt_walk g Ts) @
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   160
            (if is_countable (type_of1 (Ts, g)) andalso loose_bvar1 (g, n)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   161
                andalso indep g n 0 andalso g <> Bound n
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   162
              then [(f $ Bound (n + 1), incr_boundvars (~ n) g)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   163
              else [])
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   164
          end
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   165
        | cnt_walk _ _ = []
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   166
    in map (fn (t1, t2) => let
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   167
        val T1 = type_of1 ([T], t2)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   168
        val T2 = type_of1 ([T], t)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   169
      in ([SOME (Abs (n, T1, Abs (n, T, t1))), NONE, NONE, SOME (Abs (n, T, t2))],
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   170
        [SOME T1, SOME T, SOME T2])
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   171
      end) (cnt_walk t [T])
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   172
    end
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   173
  | cnt_prefixes _ _ = []
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   174
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   175
fun apply_dests thm dests =
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   176
  let
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   177
    fun apply thm th' =
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   178
      let
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   179
        val th'' = thm RS th'
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   180
      in [th''] @ loop th'' end
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   181
      handle (THM _) => []
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   182
    and loop thm =
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   183
      flat (map (apply thm) dests)
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   184
  in
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   185
    [thm] @ ([thm RS @{thm measurable_compose_rev}] handle (THM _) => []) @ loop thm
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   186
  end
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   187
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   188
fun prepare_facts ctxt facts = 
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   189
  let
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   190
    val dests = get_dest (Context.Proof ctxt)
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   191
    fun prep_dest thm =
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   192
      (if is_too_generic thm then [] else apply_dests thm dests) ;
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   193
    val preprocessors = (("std", prep_dest #> pair) :: get_preprocessors ctxt) ;
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   194
    fun preprocess_thm (thm, raw) =
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   195
      if raw then pair [thm] else fold_map (fn (_, proc) => proc thm) preprocessors #>> flat
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   196
    
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   197
    fun sel lv (th, (raw, lv')) = if lv = lv' then SOME (th, raw) else NONE ;
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   198
    fun get lv = ctxt |> Context.Proof |> get_thms |> rev |> map_filter (sel lv) ;
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   199
    val pre_thms = map (Simplifier.norm_hhf ctxt #> rpair false) facts @ get Concrete @ get Generic
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   200
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   201
    val (thms, ctxt) = fold_map preprocess_thm pre_thms ctxt |>> flat
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   202
  in (thms, ctxt) end
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   203
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
   204
fun measurable_tac ctxt facts =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50387
diff changeset
   205
  let
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   206
    fun debug_fact msg thm () =
61877
276ad4354069 renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
wenzelm
parents: 61424
diff changeset
   207
      msg ^ " " ^ Pretty.unformatted_string_of (Syntax.pretty_term ctxt (Thm.prop_of thm))
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   208
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   209
    fun IF' c t i = COND (c i) (t i) no_tac
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   210
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   211
    fun r_tac msg =
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   212
      if Config.get ctxt debug
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   213
      then FIRST' o
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59353
diff changeset
   214
        map (fn thm => resolve_tac ctxt [thm]
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   215
          THEN' K (debug_tac ctxt (debug_fact (msg ^ " resolved using") thm) all_tac))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59353
diff changeset
   216
      else resolve_tac ctxt
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   217
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   218
    val elem_congI = @{lemma "A = B \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A" by simp}
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   219
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59048
diff changeset
   220
    val (thms, ctxt) = prepare_facts ctxt facts
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   221
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   222
    fun is_sets_eq (Const (\<^const_name>\<open>HOL.eq\<close>, _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   223
          (Const (\<^const_name>\<open>sets\<close>, _) $ _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   224
          (Const (\<^const_name>\<open>sets\<close>, _) $ _)) = true
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   225
      | is_sets_eq (Const (\<^const_name>\<open>HOL.eq\<close>, _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   226
          (Const (\<^const_name>\<open>measurable\<close>, _) $ _ $ _) $
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63627
diff changeset
   227
          (Const (\<^const_name>\<open>measurable\<close>, _) $ _ $ _)) = true
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   228
      | is_sets_eq _ = false
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   229
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   230
    val cong_thms = get_cong (Context.Proof ctxt) @
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   231
      filter (fn thm => Thm.concl_of thm |> HOLogic.dest_Trueprop |> is_sets_eq handle TERM _ => false) facts
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   232
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   233
    fun sets_cong_tac i =
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   234
      Subgoal.FOCUS (fn {context = ctxt', prems = prems, ...} => (
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   235
        let
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   236
          val ctxt'' = Simplifier.add_prems prems ctxt'
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   237
        in
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   238
          r_tac "cong intro" [elem_congI]
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   239
          THEN' SOLVED' (fn i => REPEAT_DETERM (
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   240
              ((r_tac "cong solve" (cong_thms @ [@{thm refl}])
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   241
                ORELSE' IF' (fn i => fn thm => Thm.nprems_of thm > i)
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   242
                  (SOLVED' (asm_full_simp_tac ctxt''))) i)))
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   243
        end) 1) ctxt i
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   244
        THEN flexflex_tac ctxt
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   245
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   246
    val simp_solver_tac = 
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   247
      IF' not_measurable_prop (debug_tac ctxt (K "simp ") o SOLVED' (asm_full_simp_tac ctxt))
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   248
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   249
    val split_countable_tac =
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   250
      Subgoal.FOCUS (fn {context = ctxt, ...} => SUBGOAL (fn (t, i) =>
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   251
        let
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   252
          val f = dest_measurable_fun (HOLogic.dest_Trueprop t)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   253
          fun inst (ts, Ts) =
60807
wenzelm
parents: 60801
diff changeset
   254
            Thm.instantiate'
wenzelm
parents: 60801
diff changeset
   255
              (map (Option.map (Thm.ctyp_of ctxt)) Ts)
wenzelm
parents: 60801
diff changeset
   256
              (map (Option.map (Thm.cterm_of ctxt)) ts)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59353
diff changeset
   257
              @{thm measurable_compose_countable}
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60807
diff changeset
   258
        in r_tac "case_prod countable" (cnt_prefixes ctxt f |> map inst) i end
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   259
        handle TERM _ => no_tac) 1)
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   260
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   261
    val splitter = if Config.get ctxt split then split_countable_tac ctxt else K no_tac
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   262
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   263
    val single_step_tac =
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   264
      simp_solver_tac
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   265
      ORELSE' r_tac "step" thms
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   266
      ORELSE' splitter
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   267
      ORELSE' (CHANGED o sets_cong_tac)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   268
      ORELSE' (K (debug_tac ctxt (K "backtrack") no_tac))
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   269
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   270
  in debug_tac ctxt (K "start") (REPEAT (single_step_tac 1)) end;
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   271
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50387
diff changeset
   272
fun simproc ctxt redex =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 50387
diff changeset
   273
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   274
    val t = HOLogic.mk_Trueprop (Thm.term_of redex);
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   275
    fun tac {context = ctxt, prems = _ } =
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 58965
diff changeset
   276
      SOLVE (measurable_tac ctxt (Simplifier.prems_of ctxt));
73551
53c148e39819 proper treatment of nested antiquotations;
wenzelm
parents: 69597
diff changeset
   277
  in \<^try>\<open>Goal.prove ctxt [] [] t tac RS @{thm Eq_TrueI}\<close> end;
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   278
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   279
end
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   280