src/HOL/Lambda/InductTermi.thy
author wenzelm
Thu, 31 Aug 2000 01:42:23 +0200
changeset 9762 66f7eefb3967
parent 9716 9be481b4bc85
child 9771 54c6a2c6e569
permissions -rw-r--r--
ported HOL/Lambda/ListBeta;
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
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     6
Inductive characterization of terminating lambda terms.
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     7
Goes back to
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     8
Raamsdonk & Severi. On normalization. CWI TR CS-R9545, 1995.
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     9
Also rediscovered by Matthes and Joachimski.
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    10
*)
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    11
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    12
theory InductTermi = ListBeta:
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    13
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    14
consts
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    15
  IT :: "dB set"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    16
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    17
inductive IT
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    18
  intros
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    19
    Var: "rs : lists IT ==> Var n $$ rs : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    20
    Lambda: "r : IT ==> Abs r : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    21
    Beta: "[| (r[s/0])$$ss : IT; s : IT |] ==> (Abs r $ s)$$ss : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    22
  monos lists_mono       (* FIXME move to HOL!? *)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    23
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    24
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    25
text {* Every term in IT terminates. *}
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    26
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    27
lemma double_induction_lemma [rulify]:
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    28
  "s : termi beta ==> \<forall>t. t : termi beta -->
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    29
    (\<forall>r ss. t = r[s/0] $$ ss --> Abs r $ s $$ ss : termi beta)"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    30
  apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    31
  apply (erule thin_rl)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    32
  apply (rule allI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    33
  apply (rule impI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    34
  apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    35
  apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    36
  apply (rule accI)
9762
66f7eefb3967 ported HOL/Lambda/ListBeta;
wenzelm
parents: 9716
diff changeset
    37
  apply (tactic {* safe_tac (claset () addSEs [thm "apps_betasE"]) *})  -- FIXME
9716
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    38
     apply assumption
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    39
    apply (blast intro: subst_preserves_beta apps_preserves_beta)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    40
   apply (blast intro: apps_preserves_beta2 subst_preserves_beta2 rtrancl_converseI
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    41
     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
    42
  apply (blast dest: apps_preserves_betas)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    43
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    44
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    45
lemma IT_implies_termi: "t : IT ==> t : termi beta"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    46
  apply (erule IT.induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    47
    apply (drule rev_subsetD)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    48
     apply (rule lists_mono)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    49
     apply (rule Int_lower2)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    50
    apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    51
    apply (drule lists_accD)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    52
    apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    53
    apply (rule accI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    54
    apply (blast dest: head_Var_reduction)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    55
   apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    56
   apply (rule accI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    57
   apply blast
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    58
  apply (blast intro: double_induction_lemma)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    59
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    60
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    61
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    62
text {* Every terminating term is in IT *}
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    63
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    64
declare Var_apps_neq_Abs_apps [THEN not_sym, simp]
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
lemma [simp, THEN not_sym, simp]: "Var n $$ ss ~= Abs r $ s $$ ts"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    67
  apply (simp add: foldl_Cons [symmetric] del: foldl_Cons)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    68
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    69
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    70
lemma [simp]:
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    71
  "(Abs r $ s $$ ss = Abs r' $ s' $$ ss') = (r=r' \<and> s=s' \<and> ss=ss')"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    72
  apply (simp add: foldl_Cons [symmetric] del: foldl_Cons)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    73
  done
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    74
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    75
inductive_cases [elim!]:
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    76
  "Var n $$ ss : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    77
  "Abs t : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    78
  "Abs r $ s $$ ts : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    79
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    80
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    81
theorem termi_implies_IT: "r : termi beta ==> r : IT"
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    82
  apply (erule acc_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    83
  apply (rename_tac r)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    84
  apply (erule thin_rl)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    85
  apply (erule rev_mp)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    86
  apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    87
  apply (rule_tac t = r in Apps_dB_induct)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    88
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    89
   apply (rule IT.intros)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    90
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    91
   apply (drule bspec, assumption)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    92
   apply (erule mp)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    93
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    94
   apply (drule converseI)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    95
   apply (drule ex_step1I, assumption)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    96
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    97
   apply (rename_tac us)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    98
   apply (erule_tac x = "Var n $$ us" in allE)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
    99
   apply force
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   100
   apply (rename_tac u ts)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   101
   apply (case_tac ts)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   102
    apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   103
    apply (blast intro: IT.intros)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   104
   apply (rename_tac s ss)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   105
   apply simp
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   106
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   107
   apply (rule IT.intros)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   108
    apply (blast intro: apps_preserves_beta)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   109
   apply (erule mp)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   110
   apply clarify
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   111
   apply (rename_tac t)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   112
   apply (erule_tac x = "Abs u $ t $$ ss" in allE)
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   113
   apply force
9be481b4bc85 Lambda/InductTermi made new-style theory;
wenzelm
parents: 9102
diff changeset
   114
   done
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
   115
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
   116
end