src/HOL/Fun_Def.thy
author wenzelm
Sat, 18 Jan 2025 13:20:47 +0100
changeset 81913 5b9aca9b073b
parent 80322 b10f7c981df6
permissions -rw-r--r--
tuned names;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 54407
diff changeset
     1
(*  Title:      HOL/Fun_Def.thy
20324
71d63a30cc96 removed True_implies (cf. True_implies_equals);
wenzelm
parents: 20270
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
     3
*)
20324
71d63a30cc96 removed True_implies (cf. True_implies_equals);
wenzelm
parents: 20270
diff changeset
     4
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
     5
section \<open>Function Definitions and Termination Proofs\<close>
20324
71d63a30cc96 removed True_implies (cf. True_implies_equals);
wenzelm
parents: 20270
diff changeset
     6
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 54407
diff changeset
     7
theory Fun_Def
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
     8
  imports Basic_BNF_LFPs Partial_Function SAT
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
     9
  keywords
69913
ca515cf61651 more specific keyword kinds;
wenzelm
parents: 69605
diff changeset
    10
    "function" "termination" :: thy_goal_defn and
ca515cf61651 more specific keyword kinds;
wenzelm
parents: 69605
diff changeset
    11
    "fun" "fun_cases" :: thy_defn
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    12
begin
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    13
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    14
subsection \<open>Definitions with default value\<close>
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    15
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    16
definition THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    17
  where "THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)"
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    18
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    19
lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    20
  by (simp add: theI' THE_default_def)
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    21
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    22
lemma THE_default1_equality: "\<exists>!x. P x \<Longrightarrow> P a \<Longrightarrow> THE_default d P = a"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    23
  by (simp add: the1_equality THE_default_def)
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    24
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    25
lemma THE_default_none: "\<not> (\<exists>!x. P x) \<Longrightarrow> THE_default d P = d"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    26
  by (simp add: THE_default_def)
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    27
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    28
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
lemma fundef_ex1_existence:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    30
  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    31
  assumes ex1: "\<exists>!y. G x y"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    32
  shows "G x (f x)"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    33
  apply (simp only: f_def)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    34
  apply (rule THE_defaultI')
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    35
  apply (rule ex1)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    36
  done
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20654
diff changeset
    37
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    38
lemma fundef_ex1_uniqueness:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    39
  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    40
  assumes ex1: "\<exists>!y. G x y"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    41
  assumes elm: "G x (h x)"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    42
  shows "h x = f x"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
    43
  by (auto simp add: f_def ex1 elm THE_default1_equality[symmetric])
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    44
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    45
lemma fundef_ex1_iff:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    46
  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    47
  assumes ex1: "\<exists>!y. G x y"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    48
  shows "(G x y) = (f x = y)"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
    49
  by (auto simp add: ex1 f_def THE_default1_equality THE_defaultI')
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
    50
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    51
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    52
lemma fundef_default_value:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    53
  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    54
  assumes graph: "\<And>x y. G x y \<Longrightarrow> D x"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    55
  assumes "\<not> D x"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    56
  shows "f x = d x"
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    57
proof -
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20654
diff changeset
    58
  have "\<not>(\<exists>y. G x y)"
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    59
  proof
21512
3786eb1b69d6 Lemma "fundef_default_value" uses predicate instead of set.
krauss
parents: 21404
diff changeset
    60
    assume "\<exists>y. G x y"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    61
    then have "D x" using graph ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    62
    with \<open>\<not> D x\<close> show False ..
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    63
  qed
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    64
  then have "\<not>(\<exists>!y. G x y)" by blast
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    65
  then show ?thesis
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    66
    unfolding f_def by (rule THE_default_none)
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    67
qed
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    68
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
    69
definition in_rel_def[simp]: "in_rel R x y \<equiv> (x, y) \<in> R"
23739
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
    70
80322
b10f7c981df6 renamed theorems
desharna
parents: 79597
diff changeset
    71
lemma wf_in_rel: "wf R \<Longrightarrow> wfp (in_rel R)"
b10f7c981df6 renamed theorems
desharna
parents: 79597
diff changeset
    72
  by (simp add: wfp_def)
23739
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
    73
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    74
ML_file \<open>Tools/Function/function_core.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    75
ML_file \<open>Tools/Function/mutual.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    76
ML_file \<open>Tools/Function/pattern_split.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    77
ML_file \<open>Tools/Function/relation.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    78
ML_file \<open>Tools/Function/function_elims.ML\<close>
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
    79
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    80
method_setup relation = \<open>
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
    81
  Args.term >> (fn t => fn ctxt => SIMPLE_METHOD' (Function_Relation.relation_infer_tac ctxt t))
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    82
\<close> "prove termination using a user-specified wellfounded relation"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
    83
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    84
ML_file \<open>Tools/Function/function.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    85
ML_file \<open>Tools/Function/pat_completeness.ML\<close>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
    86
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    87
method_setup pat_completeness = \<open>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
    88
  Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    89
\<close> "prove completeness of (co)datatype patterns"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
    90
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    91
ML_file \<open>Tools/Function/fun.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
    92
ML_file \<open>Tools/Function/induction_schema.ML\<close>
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    93
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    94
method_setup induction_schema = \<open>
70520
11d8517d9384 clarified modules;
wenzelm
parents: 69913
diff changeset
    95
  Scan.succeed (CONTEXT_TACTIC oo Induction_Schema.induction_schema_tac)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    96
\<close> "prove an induction principle"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
    97
56643
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
    98
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
    99
subsection \<open>Measure functions\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   100
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   101
inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   102
  where is_measure_trivial: "is_measure f"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   103
57959
1bfed12a7646 updated to named_theorems;
wenzelm
parents: 56846
diff changeset
   104
named_theorems measure_function "rules that guide the heuristic generation of measure functions"
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
   105
ML_file \<open>Tools/Function/measure_functions.ML\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   106
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   107
lemma measure_size[measure_function]: "is_measure size"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   108
  by (rule is_measure_trivial)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   109
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   110
lemma measure_fst[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (fst p))"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   111
  by (rule is_measure_trivial)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   112
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   113
lemma measure_snd[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (snd p))"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   114
  by (rule is_measure_trivial)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   115
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
   116
ML_file \<open>Tools/Function/lexicographic_order.ML\<close>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   117
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   118
method_setup lexicographic_order = \<open>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   119
  Method.sections clasimp_modifiers >>
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   120
  (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false))
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   121
\<close> "termination prover for lexicographic orderings"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   122
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   123
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   124
subsection \<open>Congruence rules\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   125
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   126
lemma let_cong [fundef_cong]: "M = N \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x = g x) \<Longrightarrow> Let M f = Let N g"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   127
  unfolding Let_def by blast
22622
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22325
diff changeset
   128
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   129
lemmas [fundef_cong] =
67487
b4e65cd6974a drop redundant fundef_cong rule
Lars Hupel <lars.hupel@mytum.de>
parents: 67443
diff changeset
   130
  if_cong image_cong
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 55085
diff changeset
   131
  bex_cong ball_cong imp_cong map_option_cong Option.bind_cong
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   132
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   133
lemma split_cong [fundef_cong]:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   134
  "(\<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y) \<Longrightarrow> p = q \<Longrightarrow> case_prod f p = case_prod g q"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   135
  by (auto simp: split_def)
19934
8190655ea2d4 Added split_cong rule
krauss
parents: 19770
diff changeset
   136
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   137
lemma comp_cong [fundef_cong]: "f (g x) = f' (g' x') \<Longrightarrow> (f \<circ> g) x = (f' \<circ> g') x'"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   138
  by (simp only: o_apply)
19934
8190655ea2d4 Added split_cong rule
krauss
parents: 19770
diff changeset
   139
56643
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   140
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   141
subsection \<open>Simp rules for termination proofs\<close>
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   142
56643
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   143
declare
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   144
  trans_less_add1[termination_simp]
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   145
  trans_less_add2[termination_simp]
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   146
  trans_le_add1[termination_simp]
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   147
  trans_le_add2[termination_simp]
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   148
  less_imp_le_nat[termination_simp]
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   149
  le_imp_less_Suc[termination_simp]
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   150
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   151
lemma size_prod_simp[termination_simp]: "size_prod f g p = f (fst p) + g (snd p) + Suc 0"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   152
  by (induct p) auto
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   153
56643
41d3596d8a64 move size hooks together, with new one preceding old one and sharing same theory data
blanchet
parents: 56248
diff changeset
   154
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   155
subsection \<open>Decomposition\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   156
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   157
lemma less_by_empty: "A = {} \<Longrightarrow> A \<subseteq> B"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   158
  and union_comp_emptyL: "A O C = {} \<Longrightarrow> B O C = {} \<Longrightarrow> (A \<union> B) O C = {}"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   159
  and union_comp_emptyR: "A O B = {} \<Longrightarrow> A O C = {} \<Longrightarrow> A O (B \<union> C) = {}"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   160
  and wf_no_loop: "R O R = {} \<Longrightarrow> wf R"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   161
  by (auto simp add: wf_comp_self [of R])
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   162
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   163
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   164
subsection \<open>Reduction pairs\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   165
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   166
definition "reduction_pair P \<longleftrightarrow> wf (fst P) \<and> fst P O snd P \<subseteq> fst P"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   167
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   168
lemma reduction_pairI[intro]: "wf R \<Longrightarrow> R O S \<subseteq> R \<Longrightarrow> reduction_pair (R, S)"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   169
  by (auto simp: reduction_pair_def)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   170
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   171
lemma reduction_pair_lemma:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   172
  assumes rp: "reduction_pair P"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   173
  assumes "R \<subseteq> fst P"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   174
  assumes "S \<subseteq> snd P"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   175
  assumes "wf S"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   176
  shows "wf (R \<union> S)"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   177
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   178
  from rp \<open>S \<subseteq> snd P\<close> have "wf (fst P)" "fst P O S \<subseteq> fst P"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   179
    unfolding reduction_pair_def by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   180
  with \<open>wf S\<close> have "wf (fst P \<union> S)"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   181
    by (auto intro: wf_union_compatible)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   182
  moreover from \<open>R \<subseteq> fst P\<close> have "R \<union> S \<subseteq> fst P \<union> S" by auto
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   183
  ultimately show ?thesis by (rule wf_subset)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   184
qed
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   185
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   186
definition "rp_inv_image = (\<lambda>(R,S) f. (inv_image R f, inv_image S f))"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   187
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   188
lemma rp_inv_image_rp: "reduction_pair P \<Longrightarrow> reduction_pair (rp_inv_image P f)"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   189
  unfolding reduction_pair_def rp_inv_image_def split_def by force
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   190
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   191
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   192
subsection \<open>Concrete orders for SCNP termination proofs\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   193
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   194
definition "pair_less = less_than <*lex*> less_than"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67487
diff changeset
   195
definition "pair_leq = pair_less\<^sup>="
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   196
definition "max_strict = max_ext pair_less"
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   197
definition "max_weak = max_ext pair_leq \<union> {({}, {})}"
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   198
definition "min_strict = min_ext pair_less"
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   199
definition "min_weak = min_ext pair_leq \<union> {({}, {})}"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   200
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   201
lemma wf_pair_less[simp]: "wf pair_less"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   202
  by (auto simp: pair_less_def)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   203
71404
f2b783abfbe7 A few lemmas connected with orderings
paulson <lp15@cam.ac.uk>
parents: 70520
diff changeset
   204
lemma total_pair_less [iff]: "total_on A pair_less" and trans_pair_less [iff]: "trans pair_less"
72184
881bd98bddee reversing all the lex crap
paulson <lp15@cam.ac.uk>
parents: 72164
diff changeset
   205
  by (auto simp: total_on_def pair_less_def)
71404
f2b783abfbe7 A few lemmas connected with orderings
paulson <lp15@cam.ac.uk>
parents: 70520
diff changeset
   206
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61032
diff changeset
   207
text \<open>Introduction rules for \<open>pair_less\<close>/\<open>pair_leq\<close>\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   208
lemma pair_leqI1: "a < b \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   209
  and pair_leqI2: "a \<le> b \<Longrightarrow> s \<le> t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   210
  and pair_lessI1: "a < b  \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   211
  and pair_lessI2: "a \<le> b \<Longrightarrow> s < t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   212
  by (auto simp: pair_leq_def pair_less_def)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   213
72097
496cfe488d72 a few more lemmas
paulson <lp15@cam.ac.uk>
parents: 71404
diff changeset
   214
lemma pair_less_iff1 [simp]: "((x,y), (x,z)) \<in> pair_less \<longleftrightarrow> y<z"
496cfe488d72 a few more lemmas
paulson <lp15@cam.ac.uk>
parents: 71404
diff changeset
   215
  by (simp add: pair_less_def)
496cfe488d72 a few more lemmas
paulson <lp15@cam.ac.uk>
parents: 71404
diff changeset
   216
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   217
text \<open>Introduction rules for max\<close>
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   218
lemma smax_emptyI: "finite Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> ({}, Y) \<in> max_strict"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   219
  and smax_insertI:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   220
    "y \<in> Y \<Longrightarrow> (x, y) \<in> pair_less \<Longrightarrow> (X, Y) \<in> max_strict \<Longrightarrow> (insert x X, Y) \<in> max_strict"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   221
  and wmax_emptyI: "finite X \<Longrightarrow> ({}, X) \<in> max_weak"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   222
  and wmax_insertI:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   223
    "y \<in> YS \<Longrightarrow> (x, y) \<in> pair_leq \<Longrightarrow> (XS, YS) \<in> max_weak \<Longrightarrow> (insert x XS, YS) \<in> max_weak"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   224
  by (auto simp: max_strict_def max_weak_def elim!: max_ext.cases)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   225
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   226
text \<open>Introduction rules for min\<close>
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   227
lemma smin_emptyI: "X \<noteq> {} \<Longrightarrow> (X, {}) \<in> min_strict"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   228
  and smin_insertI:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   229
    "x \<in> XS \<Longrightarrow> (x, y) \<in> pair_less \<Longrightarrow> (XS, YS) \<in> min_strict \<Longrightarrow> (XS, insert y YS) \<in> min_strict"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   230
  and wmin_emptyI: "(X, {}) \<in> min_weak"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   231
  and wmin_insertI:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   232
    "x \<in> XS \<Longrightarrow> (x, y) \<in> pair_leq \<Longrightarrow> (XS, YS) \<in> min_weak \<Longrightarrow> (XS, insert y YS) \<in> min_weak"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   233
  by (auto simp: min_strict_def min_weak_def min_ext_def)
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   234
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   235
text \<open>Reduction Pairs.\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   236
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   237
lemma max_ext_compat:
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   238
  assumes "R O S \<subseteq> R"
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   239
  shows "max_ext R O (max_ext S \<union> {({}, {})}) \<subseteq> max_ext R"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   240
proof -
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   241
  have "\<And>X Y Z. (X, Y) \<in> max_ext R \<Longrightarrow> (Y, Z) \<in> max_ext S \<Longrightarrow> (X, Z) \<in> max_ext R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   242
  proof -
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   243
    fix X Y Z
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   244
    assume "(X,Y)\<in>max_ext R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   245
      "(Y, Z)\<in>max_ext S"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   246
    then have *: "finite X" "finite Y" "finite Z" "Y\<noteq>{}" "Z\<noteq>{}"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   247
      "(\<And>x. x\<in>X \<Longrightarrow> \<exists>y\<in>Y. (x, y)\<in>R)"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   248
      "(\<And>y. y\<in>Y \<Longrightarrow> \<exists>z\<in>Z. (y, z)\<in>S)"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   249
      by (auto elim: max_ext.cases)
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   250
    moreover have "\<And>x. x\<in>X \<Longrightarrow> \<exists>z\<in>Z. (x, z)\<in>R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   251
    proof -
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   252
      fix x
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   253
      assume "x\<in>X"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   254
      then obtain y where 1: "y\<in>Y" "(x, y)\<in>R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   255
        using * by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   256
      then obtain z where "z\<in>Z" "(y, z)\<in>S"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   257
        using * by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   258
      then show "\<exists>z\<in>Z. (x, z)\<in>R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   259
        using assms 1 by (auto elim: max_ext.cases)
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   260
    qed
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   261
    ultimately show "(X,Z)\<in>max_ext R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   262
      by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   263
  qed
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   264
  then show "max_ext R O (max_ext S \<union> {({}, {})}) \<subseteq> max_ext R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   265
    by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   266
qed
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   267
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   268
lemma max_rpair_set: "reduction_pair (max_strict, max_weak)"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   269
  unfolding max_strict_def max_weak_def
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   270
  apply (intro reduction_pairI max_ext_wf)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   271
   apply simp
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   272
  apply (rule max_ext_compat)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   273
  apply (auto simp: pair_less_def pair_leq_def)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   274
  done
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   275
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   276
lemma min_ext_compat:
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   277
  assumes "R O S \<subseteq> R"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   278
  shows "min_ext R O (min_ext S \<union> {({},{})}) \<subseteq> min_ext R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   279
proof -
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   280
  have "\<And>X Y Z z. \<forall>y\<in>Y. \<exists>x\<in>X. (x, y) \<in> R \<Longrightarrow> \<forall>z\<in>Z. \<exists>y\<in>Y. (y, z) \<in> S
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   281
  \<Longrightarrow> z \<in> Z \<Longrightarrow> \<exists>x\<in>X. (x, z) \<in> R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   282
  proof -
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   283
    fix X Y Z z
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   284
    assume *: "\<forall>y\<in>Y. \<exists>x\<in>X. (x, y) \<in> R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   285
      "\<forall>z\<in>Z. \<exists>y\<in>Y. (y, z) \<in> S"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   286
      "z\<in>Z"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   287
    then obtain y' where 1: "y'\<in>Y" "(y', z) \<in> S"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   288
      by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   289
    then obtain x' where 2: "x'\<in>X" "(x', y') \<in> R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   290
      using * by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   291
    show "\<exists>x\<in>X. (x, z) \<in> R"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   292
      using 1 2 assms by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   293
  qed
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   294
  then show ?thesis
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   295
    using assms by (auto simp: min_ext_def)
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 72184
diff changeset
   296
qed
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   297
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   298
lemma min_rpair_set: "reduction_pair (min_strict, min_weak)"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47432
diff changeset
   299
  unfolding min_strict_def min_weak_def
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   300
  apply (intro reduction_pairI min_ext_wf)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   301
   apply simp
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   302
  apply (rule min_ext_compat)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   303
  apply (auto simp: pair_less_def pair_leq_def)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 63432
diff changeset
   304
  done
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   305
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   306
79597
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   307
subsection \<open>Yet more induction principles on the natural numbers\<close>
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   308
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   309
lemma nat_descend_induct [case_names base descend]:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   310
  fixes P :: "nat \<Rightarrow> bool"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   311
  assumes H1: "\<And>k. k > n \<Longrightarrow> P k"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   312
  assumes H2: "\<And>k. k \<le> n \<Longrightarrow> (\<And>i. i > k \<Longrightarrow> P i) \<Longrightarrow> P k"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   313
  shows "P m"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   314
  using assms by induction_schema (force intro!: wf_measure [of "\<lambda>k. Suc n - k"])+
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   315
79597
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   316
lemma induct_nat_012[case_names 0 1 ge2]:
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   317
  "P 0 \<Longrightarrow> P (Suc 0) \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (Suc n) \<Longrightarrow> P (Suc (Suc n))) \<Longrightarrow> P n"
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   318
proof (induction_schema, pat_completeness)
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   319
  show "wf (Wellfounded.measure id)"
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   320
    by simp
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   321
qed auto
76a1c0ea6777 A few lemmas brought in from AFP entries
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   322
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 63654
diff changeset
   323
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   324
subsection \<open>Tool setup\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   325
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
   326
ML_file \<open>Tools/Function/termination.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
   327
ML_file \<open>Tools/Function/scnp_solve.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
   328
ML_file \<open>Tools/Function/scnp_reconstruct.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67613
diff changeset
   329
ML_file \<open>Tools/Function/fun_cases.ML\<close>
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   330
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 64591
diff changeset
   331
ML_val \<comment> \<open>setup inactive\<close>
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   332
\<open>
36521
73ed9f18fdd3 default termination prover as plain tactic
krauss
parents: 34228
diff changeset
   333
  Context.theory_map (Function_Common.set_termination_prover
60682
5a6cd2560549 have the installed termination prover take a 'quiet' flag
blanchet
parents: 59953
diff changeset
   334
    (K (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])))
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60682
diff changeset
   335
\<close>
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   336
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   337
end