src/HOL/FunDef.thy
author huffman
Wed, 18 Apr 2012 15:48:32 +0200
changeset 47544 e455cdaac479
parent 47432 e1576d13e933
child 47701 157e6108a342
permissions -rw-r--r--
move constant 'Respects' into Lifting.thy; add quantifier transfer rules for quotients
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20324
71d63a30cc96 removed True_implies (cf. True_implies_equals);
wenzelm
parents: 20270
diff changeset
     1
(*  Title:      HOL/FunDef.thy
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
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
     5
header {* Function Definitions and Termination Proofs *}
20324
71d63a30cc96 removed True_implies (cf. True_implies_equals);
wenzelm
parents: 20270
diff changeset
     6
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
     7
theory FunDef
40108
dbab949c2717 integrated partial_function into HOL-Plain
krauss
parents: 37767
diff changeset
     8
imports Partial_Function Wellfounded
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46526
diff changeset
     9
keywords "function" "termination" :: thy_goal and "fun" :: thy_decl
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    10
uses
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
    11
  "Tools/prop_logic.ML"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
    12
  "Tools/sat_solver.ML"
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
    13
  ("Tools/Function/function_common.ML")
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    14
  ("Tools/Function/context_tree.ML")
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
    15
  ("Tools/Function/function_core.ML")
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    16
  ("Tools/Function/sum_tree.ML")
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    17
  ("Tools/Function/mutual.ML")
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    18
  ("Tools/Function/pattern_split.ML")
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
    19
  ("Tools/Function/function.ML")
33100
acc2bd3934ba renamed auto_term.ML -> relation.ML
krauss
parents: 33099
diff changeset
    20
  ("Tools/Function/relation.ML")
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    21
  ("Tools/Function/measure_functions.ML")
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    22
  ("Tools/Function/lexicographic_order.ML")
33083
1fad3160d873 pat_completeness gets its own file
krauss
parents: 32235
diff changeset
    23
  ("Tools/Function/pat_completeness.ML")
33098
3e9ae9032273 renamed FundefDatatype -> Function_Fun
krauss
parents: 33083
diff changeset
    24
  ("Tools/Function/fun.ML")
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents: 33348
diff changeset
    25
  ("Tools/Function/induction_schema.ML")
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    26
  ("Tools/Function/termination.ML")
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    27
  ("Tools/Function/scnp_solve.ML")
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
    28
  ("Tools/Function/scnp_reconstruct.ML")
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    29
begin
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    30
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
    31
subsection {* Definitions with default value. *}
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    32
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    33
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21364
diff changeset
    34
  THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    35
  "THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)"
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    36
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    37
lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    38
  by (simp add: theI' THE_default_def)
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    39
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    40
lemma THE_default1_equality:
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    41
    "\<lbrakk>\<exists>!x. P x; P a\<rbrakk> \<Longrightarrow> THE_default d P = a"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    42
  by (simp add: the1_equality THE_default_def)
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    43
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    44
lemma THE_default_none:
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    45
    "\<not>(\<exists>!x. P x) \<Longrightarrow> THE_default d P = d"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    46
  by (simp add:THE_default_def)
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    47
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    48
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    49
lemma fundef_ex1_existence:
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    50
  assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    51
  assumes ex1: "\<exists>!y. G x y"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    52
  shows "G x (f x)"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    53
  apply (simp only: f_def)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    54
  apply (rule THE_defaultI')
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    55
  apply (rule ex1)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    56
  done
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20654
diff changeset
    57
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    58
lemma fundef_ex1_uniqueness:
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    59
  assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    60
  assumes ex1: "\<exists>!y. G x y"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    61
  assumes elm: "G x (h x)"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    62
  shows "h x = f x"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    63
  apply (simp only: f_def)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    64
  apply (rule THE_default1_equality [symmetric])
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    65
   apply (rule ex1)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    66
  apply (rule elm)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    67
  done
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    68
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    69
lemma fundef_ex1_iff:
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    70
  assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    71
  assumes ex1: "\<exists>!y. G x y"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    72
  shows "(G x y) = (f x = y)"
20536
f088edff8af8 Function package: Outside their domain functions now return "arbitrary".
krauss
parents: 20523
diff changeset
    73
  apply (auto simp:ex1 f_def THE_default1_equality)
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    74
  apply (rule THE_defaultI')
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    75
  apply (rule ex1)
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    76
  done
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
    77
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    78
lemma fundef_default_value:
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    79
  assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    80
  assumes graph: "\<And>x y. G x y \<Longrightarrow> D x"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    81
  assumes "\<not> D x"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    82
  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
    83
proof -
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20654
diff changeset
    84
  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
    85
  proof
21512
3786eb1b69d6 Lemma "fundef_default_value" uses predicate instead of set.
krauss
parents: 21404
diff changeset
    86
    assume "\<exists>y. G x y"
3786eb1b69d6 Lemma "fundef_default_value" uses predicate instead of set.
krauss
parents: 21404
diff changeset
    87
    hence "D x" using graph ..
3786eb1b69d6 Lemma "fundef_default_value" uses predicate instead of set.
krauss
parents: 21404
diff changeset
    88
    with `\<not> D x` show False ..
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    89
  qed
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20654
diff changeset
    90
  hence "\<not>(\<exists>!y. G x y)" by blast
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
    91
20654
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    92
  thus ?thesis
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    93
    unfolding f_def
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    94
    by (rule THE_default_none)
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    95
qed
d80502f0d701 1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents: 20536
diff changeset
    96
23739
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
    97
definition in_rel_def[simp]:
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
    98
  "in_rel R x y == (x, y) \<in> R"
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
    99
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
   100
lemma wf_in_rel:
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
   101
  "wf R \<Longrightarrow> wfP (in_rel R)"
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
   102
  by (simp add: wfP_def)
c5ead5df7f35 Inserted definition of in_rel again (since member2 was removed).
berghofe
parents: 23494
diff changeset
   103
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
   104
use "Tools/Function/function_common.ML"
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   105
use "Tools/Function/context_tree.ML"
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
   106
use "Tools/Function/function_core.ML"
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   107
use "Tools/Function/sum_tree.ML"
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   108
use "Tools/Function/mutual.ML"
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   109
use "Tools/Function/pattern_split.ML"
33100
acc2bd3934ba renamed auto_term.ML -> relation.ML
krauss
parents: 33099
diff changeset
   110
use "Tools/Function/relation.ML"
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
   111
use "Tools/Function/function.ML"
33083
1fad3160d873 pat_completeness gets its own file
krauss
parents: 32235
diff changeset
   112
use "Tools/Function/pat_completeness.ML"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   113
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   114
method_setup pat_completeness = {*
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   115
  Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac)
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   116
*} "prove completeness of datatype patterns"
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   117
33098
3e9ae9032273 renamed FundefDatatype -> Function_Fun
krauss
parents: 33083
diff changeset
   118
use "Tools/Function/fun.ML"
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents: 33348
diff changeset
   119
use "Tools/Function/induction_schema.ML"
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   120
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   121
method_setup induction_schema = {*
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   122
  Scan.succeed (RAW_METHOD o Induction_Schema.induction_schema_tac)
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   123
*} "prove an induction principle"
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   124
25567
5720345ea689 experimental version of automated induction scheme generator (cf. HOL/ex/Induction_Scheme.thy)
krauss
parents: 25556
diff changeset
   125
setup {* 
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
   126
  Function.setup
33098
3e9ae9032273 renamed FundefDatatype -> Function_Fun
krauss
parents: 33083
diff changeset
   127
  #> Function_Fun.setup
25567
5720345ea689 experimental version of automated induction scheme generator (cf. HOL/ex/Induction_Scheme.thy)
krauss
parents: 25556
diff changeset
   128
*}
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents: 19564
diff changeset
   129
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   130
subsection {* Measure Functions *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   131
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   132
inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   133
where is_measure_trivial: "is_measure f"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   134
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   135
use "Tools/Function/measure_functions.ML"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   136
setup MeasureFunctions.setup
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   137
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   138
lemma measure_size[measure_function]: "is_measure size"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   139
by (rule is_measure_trivial)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   140
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   141
lemma measure_fst[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (fst p))"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   142
by (rule is_measure_trivial)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   143
lemma measure_snd[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (snd p))"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   144
by (rule is_measure_trivial)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   145
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   146
use "Tools/Function/lexicographic_order.ML"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   147
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   148
method_setup lexicographic_order = {*
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   149
  Method.sections clasimp_modifiers >>
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   150
  (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false))
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   151
*} "termination prover for lexicographic orderings"
e1576d13e933 more standard method setup;
wenzelm
parents: 46950
diff changeset
   152
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents: 33098
diff changeset
   153
setup Lexicographic_Order.setup 
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   154
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   155
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   156
subsection {* Congruence Rules *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   157
22838
haftmann
parents: 22816
diff changeset
   158
lemma let_cong [fundef_cong]:
haftmann
parents: 22816
diff changeset
   159
  "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
   160
  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
   161
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   162
lemmas [fundef_cong] =
22838
haftmann
parents: 22816
diff changeset
   163
  if_cong image_cong INT_cong UN_cong
46526
c4cf9d03c352 added congruence rules for Option.{map|bind}
krauss
parents: 40108
diff changeset
   164
  bex_cong ball_cong imp_cong Option.map_cong Option.bind_cong
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   165
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   166
lemma split_cong [fundef_cong]:
22838
haftmann
parents: 22816
diff changeset
   167
  "(\<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y) \<Longrightarrow> p = q
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   168
    \<Longrightarrow> split f p = split g q"
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   169
  by (auto simp: split_def)
19934
8190655ea2d4 Added split_cong rule
krauss
parents: 19770
diff changeset
   170
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   171
lemma comp_cong [fundef_cong]:
22838
haftmann
parents: 22816
diff changeset
   172
  "f (g x) = f' (g' x') \<Longrightarrow> (f o g) x = (f' o g') x'"
22816
0eba117368d9 added header;
wenzelm
parents: 22622
diff changeset
   173
  unfolding o_apply .
19934
8190655ea2d4 Added split_cong rule
krauss
parents: 19770
diff changeset
   174
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   175
subsection {* Simp rules for termination proofs *}
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   176
26749
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26748
diff changeset
   177
lemma termination_basic_simps[termination_simp]:
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26748
diff changeset
   178
  "x < (y::nat) \<Longrightarrow> x < y + z" 
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26748
diff changeset
   179
  "x < z \<Longrightarrow> x < y + z"
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   180
  "x \<le> y \<Longrightarrow> x \<le> y + (z::nat)"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   181
  "x \<le> z \<Longrightarrow> x \<le> y + (z::nat)"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   182
  "x < y \<Longrightarrow> x \<le> (y::nat)"
26749
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26748
diff changeset
   183
by arith+
397a1aeede7d * New attribute "termination_simp": Simp rules for termination proofs
krauss
parents: 26748
diff changeset
   184
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   185
declare le_imp_less_Suc[termination_simp]
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   186
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   187
lemma prod_size_simp[termination_simp]:
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   188
  "prod_size f g p = f (fst p) + g (snd p) + Suc 0"
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   189
by (induct p) auto
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   190
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   191
subsection {* Decomposition *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   192
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   193
lemma less_by_empty: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   194
  "A = {} \<Longrightarrow> A \<subseteq> B"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   195
and  union_comp_emptyL:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   196
  "\<lbrakk> A O C = {}; B O C = {} \<rbrakk> \<Longrightarrow> (A \<union> B) O C = {}"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   197
and union_comp_emptyR:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   198
  "\<lbrakk> A O B = {}; A O C = {} \<rbrakk> \<Longrightarrow> A O (B \<union> C) = {}"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   199
and wf_no_loop: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   200
  "R O R = {} \<Longrightarrow> wf R"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   201
by (auto simp add: wf_comp_self[of R])
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   202
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   203
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   204
subsection {* Reduction Pairs *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   205
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   206
definition
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   207
  "reduction_pair P = (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
   208
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   209
lemma reduction_pairI[intro]: "wf R \<Longrightarrow> R O S \<subseteq> R \<Longrightarrow> reduction_pair (R, S)"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   210
unfolding reduction_pair_def by auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   211
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   212
lemma reduction_pair_lemma:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   213
  assumes rp: "reduction_pair P"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   214
  assumes "R \<subseteq> fst P"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   215
  assumes "S \<subseteq> snd P"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   216
  assumes "wf S"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   217
  shows "wf (R \<union> S)"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   218
proof -
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   219
  from rp `S \<subseteq> snd P` 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
   220
    unfolding reduction_pair_def by auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   221
  with `wf S` have "wf (fst P \<union> S)" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   222
    by (auto intro: wf_union_compatible)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   223
  moreover from `R \<subseteq> fst P` have "R \<union> S \<subseteq> fst P \<union> S" by auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   224
  ultimately show ?thesis by (rule wf_subset) 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   225
qed
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   226
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   227
definition
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   228
  "rp_inv_image = (\<lambda>(R,S) f. (inv_image R f, inv_image S f))"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   229
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   230
lemma rp_inv_image_rp:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   231
  "reduction_pair P \<Longrightarrow> reduction_pair (rp_inv_image P f)"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   232
  unfolding reduction_pair_def rp_inv_image_def split_def
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   233
  by force
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   234
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   235
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   236
subsection {* Concrete orders for SCNP termination proofs *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   237
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   238
definition "pair_less = less_than <*lex*> less_than"
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   239
definition "pair_leq = pair_less^="
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   240
definition "max_strict = max_ext pair_less"
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   241
definition "max_weak = max_ext pair_leq \<union> {({}, {})}"
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   242
definition "min_strict = min_ext pair_less"
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36521
diff changeset
   243
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
   244
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   245
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
   246
  by (auto simp: pair_less_def)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   247
29127
2a684ee023e7 proper document antiquotations;
wenzelm
parents: 29125
diff changeset
   248
text {* Introduction rules for @{text pair_less}/@{text pair_leq} *}
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   249
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
   250
  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
   251
  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
   252
  and pair_lessI2: "a \<le> b \<Longrightarrow> s < t \<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
   253
  unfolding pair_leq_def pair_less_def by auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   254
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   255
text {* Introduction rules for max *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   256
lemma smax_emptyI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   257
  "finite Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> ({}, Y) \<in> max_strict" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   258
  and smax_insertI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   259
  "\<lbrakk>y \<in> Y; (x, y) \<in> pair_less; (X, Y) \<in> max_strict\<rbrakk> \<Longrightarrow> (insert x X, Y) \<in> max_strict"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   260
  and wmax_emptyI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   261
  "finite X \<Longrightarrow> ({}, X) \<in> max_weak" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   262
  and wmax_insertI:
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   263
  "\<lbrakk>y \<in> YS; (x, y) \<in> pair_leq; (XS, YS) \<in> max_weak\<rbrakk> \<Longrightarrow> (insert x XS, YS) \<in> max_weak" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   264
unfolding max_strict_def max_weak_def by (auto elim!: max_ext.cases)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   265
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   266
text {* Introduction rules for min *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   267
lemma smin_emptyI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   268
  "X \<noteq> {} \<Longrightarrow> (X, {}) \<in> min_strict" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   269
  and smin_insertI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   270
  "\<lbrakk>x \<in> XS; (x, y) \<in> pair_less; (XS, YS) \<in> min_strict\<rbrakk> \<Longrightarrow> (XS, insert y YS) \<in> min_strict"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   271
  and wmin_emptyI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   272
  "(X, {}) \<in> min_weak" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   273
  and wmin_insertI: 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   274
  "\<lbrakk>x \<in> XS; (x, y) \<in> pair_leq; (XS, YS) \<in> min_weak\<rbrakk> \<Longrightarrow> (XS, insert y YS) \<in> min_weak" 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   275
by (auto simp: min_strict_def min_weak_def min_ext_def)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   276
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   277
text {* Reduction Pairs *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   278
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   279
lemma max_ext_compat: 
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   280
  assumes "R O S \<subseteq> R"
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   281
  shows "max_ext R O (max_ext S \<union> {({},{})}) \<subseteq> max_ext R"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   282
using assms 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   283
apply auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   284
apply (elim max_ext.cases)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   285
apply rule
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   286
apply auto[3]
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   287
apply (drule_tac x=xa in meta_spec)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   288
apply simp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   289
apply (erule bexE)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   290
apply (drule_tac x=xb in meta_spec)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   291
by auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   292
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   293
lemma max_rpair_set: "reduction_pair (max_strict, max_weak)"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   294
  unfolding max_strict_def max_weak_def 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   295
apply (intro reduction_pairI max_ext_wf)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   296
apply simp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   297
apply (rule max_ext_compat)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   298
by (auto simp: pair_less_def pair_leq_def)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   299
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   300
lemma min_ext_compat: 
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   301
  assumes "R O S \<subseteq> R"
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 31775
diff changeset
   302
  shows "min_ext R O  (min_ext S \<union> {({},{})}) \<subseteq> min_ext R"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   303
using assms 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   304
apply (auto simp: min_ext_def)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   305
apply (drule_tac x=ya in bspec, assumption)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   306
apply (erule bexE)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   307
apply (drule_tac x=xc in bspec)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   308
apply assumption
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   309
by auto
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   310
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   311
lemma min_rpair_set: "reduction_pair (min_strict, min_weak)"
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   312
  unfolding min_strict_def min_weak_def 
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   313
apply (intro reduction_pairI min_ext_wf)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   314
apply simp
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   315
apply (rule min_ext_compat)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   316
by (auto simp: pair_less_def pair_leq_def)
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   317
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   318
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   319
subsection {* Tool setup *}
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   320
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   321
use "Tools/Function/termination.ML"
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   322
use "Tools/Function/scnp_solve.ML"
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 31723
diff changeset
   323
use "Tools/Function/scnp_reconstruct.ML"
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   324
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   325
setup {* ScnpReconstruct.setup *}
30480
f3421e8379ab keep dead code fresh;
wenzelm
parents: 30446
diff changeset
   326
f3421e8379ab keep dead code fresh;
wenzelm
parents: 30446
diff changeset
   327
ML_val -- "setup inactive"
f3421e8379ab keep dead code fresh;
wenzelm
parents: 30446
diff changeset
   328
{*
36521
73ed9f18fdd3 default termination prover as plain tactic
krauss
parents: 34228
diff changeset
   329
  Context.theory_map (Function_Common.set_termination_prover
73ed9f18fdd3 default termination prover as plain tactic
krauss
parents: 34228
diff changeset
   330
    (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS]))
29125
d41182a8135c method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents: 27271
diff changeset
   331
*}
26875
e18574413bc4 Measure functions can now be declared via special rules, allowing for a
krauss
parents: 26749
diff changeset
   332
19564
d3e2f532459a First usable version of the new function definition package (HOL/function_packake/...).
krauss
parents:
diff changeset
   333
end