src/HOL/Recdef.thy
author huffman
Wed, 18 Feb 2009 15:01:53 -0800
changeset 29981 7d0ed261b712
parent 29654 24e73987bfe2
child 31723 f5cafe803b55
permissions -rw-r--r--
generalize int_dvd_cancel_factor simproc to idom class
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7701
2c8c3b7003e5 load / setup recdef package (TFL);
wenzelm
parents: 7357
diff changeset
     1
(*  Title:      HOL/Recdef.thy
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
     2
    Author:     Konrad Slind and Markus Wenzel, TU Muenchen
12023
wenzelm
parents: 11533
diff changeset
     3
*)
5123
97c1d5c7b701 stepping stones;
wenzelm
parents:
diff changeset
     4
12023
wenzelm
parents: 11533
diff changeset
     5
header {* TFL: recursive function definitions *}
7701
2c8c3b7003e5 load / setup recdef package (TFL);
wenzelm
parents: 7357
diff changeset
     6
15131
c69542757a4d New theory header syntax.
nipkow
parents: 12437
diff changeset
     7
theory Recdef
29654
24e73987bfe2 Plain, Main form meeting points in import hierarchy
haftmann
parents: 26748
diff changeset
     8
imports FunDef Plain
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15481
diff changeset
     9
uses
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    10
  ("Tools/TFL/casesplit.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    11
  ("Tools/TFL/utils.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    12
  ("Tools/TFL/usyntax.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    13
  ("Tools/TFL/dcterm.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    14
  ("Tools/TFL/thms.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    15
  ("Tools/TFL/rules.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    16
  ("Tools/TFL/thry.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    17
  ("Tools/TFL/tfl.ML")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    18
  ("Tools/TFL/post.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 12437
diff changeset
    19
  ("Tools/recdef_package.ML")
c69542757a4d New theory header syntax.
nipkow
parents: 12437
diff changeset
    20
begin
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    21
26748
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    22
text{** This form avoids giant explosions in proofs.  NOTE USE OF ==*}
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    23
lemma def_wfrec: "[| f==wfrec r H;  wf(r) |] ==> f(a) = H (cut f r a) a"
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    24
apply auto
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    25
apply (blast intro: wfrec)
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    26
done
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    27
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    28
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    29
lemma tfl_wf_induct: "ALL R. wf R -->  
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    30
       (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))"
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    31
apply clarify
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    32
apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast)
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    33
done
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    34
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    35
lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)"
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    36
apply clarify
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    37
apply (rule cut_apply, assumption)
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    38
done
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    39
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    40
lemma tfl_wfrec:
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    41
     "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)"
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    42
apply clarify
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    43
apply (erule wfrec)
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    44
done
4d51ddd6aa5c Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents: 23742
diff changeset
    45
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    46
lemma tfl_eq_True: "(x = True) --> x"
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    47
  by blast
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    48
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    49
lemma tfl_rev_eq_mp: "(x = y) --> y --> x";
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    50
  by blast
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    51
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    52
lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)"
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    53
  by blast
6438
e55a1869ed38 'HOL/recdef' theory data;
wenzelm
parents: 5123
diff changeset
    54
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    55
lemma tfl_P_imp_P_iff_True: "P ==> P = True"
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    56
  by blast
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    57
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    58
lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)"
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    59
  by blast
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    60
12023
wenzelm
parents: 11533
diff changeset
    61
lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)"
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    62
  by simp
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    63
12023
wenzelm
parents: 11533
diff changeset
    64
lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R"
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    65
  by blast
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    66
12023
wenzelm
parents: 11533
diff changeset
    67
lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q"
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    68
  by blast
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    69
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    70
use "Tools/TFL/casesplit.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    71
use "Tools/TFL/utils.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    72
use "Tools/TFL/usyntax.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    73
use "Tools/TFL/dcterm.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    74
use "Tools/TFL/thms.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    75
use "Tools/TFL/rules.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    76
use "Tools/TFL/thry.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    77
use "Tools/TFL/tfl.ML"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    78
use "Tools/TFL/post.ML"
10773
0deff0197496 renamed .sml files to .ML;
wenzelm
parents: 10653
diff changeset
    79
use "Tools/recdef_package.ML"
6438
e55a1869ed38 'HOL/recdef' theory data;
wenzelm
parents: 5123
diff changeset
    80
setup RecdefPackage.setup
e55a1869ed38 'HOL/recdef' theory data;
wenzelm
parents: 5123
diff changeset
    81
9855
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    82
lemmas [recdef_simp] =
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    83
  inv_image_def
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    84
  measure_def
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    85
  lex_prod_def
11284
981ea92a86dd made same_fst recdef_simp
nipkow
parents: 11165
diff changeset
    86
  same_fst_def
9855
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    87
  less_Suc_eq [THEN iffD2]
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    88
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents: 22622
diff changeset
    89
lemmas [recdef_cong] =
22622
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22399
diff changeset
    90
  if_cong let_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong
9855
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    91
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    92
lemmas [recdef_wf] =
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    93
  wf_trancl
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    94
  wf_less_than
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    95
  wf_lex_prod
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    96
  wf_inv_image
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    97
  wf_measure
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
    98
  wf_pred_nat
10653
55f33da63366 small mods.
nipkow
parents: 10212
diff changeset
    99
  wf_same_fst
9855
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
   100
  wf_empty
709a295731e2 improved recdef setup;
wenzelm
parents: 8303
diff changeset
   101
6438
e55a1869ed38 'HOL/recdef' theory data;
wenzelm
parents: 5123
diff changeset
   102
end