src/HOL/Lambda/InductTermi.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 12011 1a3a7b3cd9bb
child 18241 afdba6b3e383
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Lambda/InductTermi.thy
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     4
    Copyright   1998 TU Muenchen
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     5
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
     6
Inductive characterization of terminating lambda terms.  Goes back to
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
     7
Raamsdonk & Severi. On normalization. CWI TR CS-R9545, 1995.  Also
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
     8
rediscovered by Matthes and Joachimski.
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     9
*)
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    10
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    11
header {* Inductive characterization of terminating lambda terms *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    12
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12011
diff changeset
    13
theory InductTermi imports ListBeta begin
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    14
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    15
subsection {* Terminating lambda terms *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    16
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    17
consts
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    18
  IT :: "dB set"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    19
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    20
inductive IT
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    21
  intros
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    22
    Var [intro]: "rs : lists IT ==> Var n \<degree>\<degree> rs : IT"
11946
wenzelm
parents: 11639
diff changeset
    23
    Lambda [intro]: "r : IT ==> Abs r : IT"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    24
    Beta [intro]: "(r[s/0]) \<degree>\<degree> ss : IT ==> s : IT ==> (Abs r \<degree> s) \<degree>\<degree> ss : IT"
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    25
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    26
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    27
subsection {* Every term in IT terminates *}
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    28
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
    29
lemma double_induction_lemma [rule_format]:
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    30
  "s : termi beta ==> \<forall>t. t : termi beta -->
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    31
    (\<forall>r ss. t = r[s/0] \<degree>\<degree> ss --> Abs r \<degree> s \<degree>\<degree> ss : termi beta)"
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    32
  apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    33
  apply (erule thin_rl)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    34
  apply (rule allI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    35
  apply (rule impI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    36
  apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    37
  apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    38
  apply (rule accI)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9762
diff changeset
    39
  apply (safe elim!: apps_betasE)
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    40
     apply assumption
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    41
    apply (blast intro: subst_preserves_beta apps_preserves_beta)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    42
   apply (blast intro: apps_preserves_beta2 subst_preserves_beta2 rtrancl_converseI
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    43
     dest: acc_downwards)  (* FIXME: acc_downwards can be replaced by acc(R ^* ) = acc(r) *)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    44
  apply (blast dest: apps_preserves_betas)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    45
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    46
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    47
lemma IT_implies_termi: "t : IT ==> t : termi beta"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    48
  apply (erule IT.induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    49
    apply (drule rev_subsetD)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    50
     apply (rule lists_mono)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    51
     apply (rule Int_lower2)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    52
    apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    53
    apply (drule lists_accD)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    54
    apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    55
    apply (rule accI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    56
    apply (blast dest: head_Var_reduction)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    57
   apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    58
   apply (rule accI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    59
   apply blast
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    60
  apply (blast intro: double_induction_lemma)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    61
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    62
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    63
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    64
subsection {* Every terminating term is in IT *}
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    65
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    66
declare Var_apps_neq_Abs_apps [THEN not_sym, simp]
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    67
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    68
lemma [simp, THEN not_sym, simp]: "Var n \<degree>\<degree> ss \<noteq> Abs r \<degree> s \<degree>\<degree> ts"
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    69
  apply (simp add: foldl_Cons [symmetric] del: foldl_Cons)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    70
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    71
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    72
lemma [simp]:
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    73
  "(Abs r \<degree> s \<degree>\<degree> ss = Abs r' \<degree> s' \<degree>\<degree> ss') = (r = r' \<and> s = s' \<and> ss = ss')"
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    74
  apply (simp add: foldl_Cons [symmetric] del: foldl_Cons)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    75
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    76
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    77
inductive_cases [elim!]:
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    78
  "Var n \<degree>\<degree> ss : IT"
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    79
  "Abs t : IT"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    80
  "Abs r \<degree> s \<degree>\<degree> ts : IT"
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    81
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    82
theorem termi_implies_IT: "r : termi beta ==> r : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    83
  apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    84
  apply (rename_tac r)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    85
  apply (erule thin_rl)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    86
  apply (erule rev_mp)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    87
  apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    88
  apply (rule_tac t = r in Apps_dB_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    89
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    90
   apply (rule IT.intros)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    91
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    92
   apply (drule bspec, assumption)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    93
   apply (erule mp)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    94
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    95
   apply (drule converseI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    96
   apply (drule ex_step1I, assumption)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    97
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    98
   apply (rename_tac us)
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
    99
   apply (erule_tac x = "Var n \<degree>\<degree> us" in allE)
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   100
   apply force
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   101
   apply (rename_tac u ts)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   102
   apply (case_tac ts)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   103
    apply simp
11946
wenzelm
parents: 11639
diff changeset
   104
    apply blast
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   105
   apply (rename_tac s ss)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   106
   apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   107
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   108
   apply (rule IT.intros)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   109
    apply (blast intro: apps_preserves_beta)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   110
   apply (erule mp)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   111
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   112
   apply (rename_tac t)
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11946
diff changeset
   113
   apply (erule_tac x = "Abs u \<degree> t \<degree>\<degree> ss" in allE)
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   114
   apply force
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   115
   done
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
   116
11639
wenzelm
parents: 9941
diff changeset
   117
end