src/Cube/ex/ex.thy
author wenzelm
Mon, 05 Sep 2005 17:38:17 +0200
changeset 17260 df7c3b1f390a
parent 17255 f9c5cbdce422
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17260
wenzelm
parents: 17255
diff changeset
     1
wenzelm
parents: 17255
diff changeset
     2
(* $Id$ *)
wenzelm
parents: 17255
diff changeset
     3
wenzelm
parents: 17255
diff changeset
     4
header {* Lambda Cube Examples *}
4583
6d9be46ea566 reorganized into individual theories;
wenzelm
parents:
diff changeset
     5
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
     6
theory ex
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
     7
imports Cube
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
     8
begin
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
     9
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    10
text {*
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    11
  Examples taken from:
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    12
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    13
  H. Barendregt. Introduction to Generalised Type Systems.
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    14
  J. Functional Programming.
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    15
*}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    16
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    17
method_setup depth_solve = {*
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    18
  Method.thms_args (fn thms => Method.METHOD (fn facts =>
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    19
  (DEPTH_SOLVE (HEADGOAL (ares_tac (facts @ thms))))))
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    20
*} ""
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    21
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    22
method_setup depth_solve1 = {*
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    23
  Method.thms_args (fn thms => Method.METHOD (fn facts =>
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    24
  (DEPTH_SOLVE_1 (HEADGOAL (ares_tac (facts @ thms))))))
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    25
*} ""
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    26
17255
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
    27
method_setup strip_asms =  {*
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
    28
  let val strip_b = thm "strip_b" and strip_s = thm "strip_s" in
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
    29
    Method.thms_args (fn thms => Method.METHOD (fn facts =>
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
    30
      REPEAT (resolve_tac [strip_b, strip_s] 1 THEN DEPTH_SOLVE_1 (ares_tac (facts @ thms) 1))))
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
    31
  end
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
    32
*} ""
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    33
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    34
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    35
subsection {* Simple types *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    36
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    37
lemma "A:* |- A->A : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    38
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    39
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    40
lemma "A:* |- Lam a:A. a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    41
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    42
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    43
lemma "A:* B:* b:B |- Lam x:A. b : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    44
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    45
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    46
lemma "A:* b:A |- (Lam a:A. a)^b: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    47
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    48
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    49
lemma "A:* B:* c:A b:B |- (Lam x:A. b)^ c: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    50
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    51
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    52
lemma "A:* B:* |- Lam a:A. Lam b:B. a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    53
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    54
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    55
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    56
subsection {* Second-order types *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    57
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    58
lemma (in L2) "|- Lam A:*. Lam a:A. a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    59
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    60
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    61
lemma (in L2) "A:* |- (Lam B:*.Lam b:B. b)^A : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    62
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    63
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    64
lemma (in L2) "A:* b:A |- (Lam B:*.Lam b:B. b) ^ A ^ b: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    65
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    66
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    67
lemma (in L2) "|- Lam B:*.Lam a:(Pi A:*.A).a ^ ((Pi A:*.A)->B) ^ a: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    68
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    69
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    70
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    71
subsection {* Weakly higher-order propositional logic *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    72
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    73
lemma (in Lomega) "|- Lam A:*.A->A : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    74
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    75
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    76
lemma (in Lomega) "B:* |- (Lam A:*.A->A) ^ B : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    77
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    78
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    79
lemma (in Lomega) "B:* b:B |- (Lam y:B. b): ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    80
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    81
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    82
lemma (in Lomega) "A:* F:*->* |- F^(F^A): ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    83
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    84
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    85
lemma (in Lomega) "A:* |- Lam F:*->*.F^(F^A): ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    86
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    87
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    88
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    89
subsection {* LP *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    90
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    91
lemma (in LP) "A:* |- A -> * : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    92
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    93
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    94
lemma (in LP) "A:* P:A->* a:A |- P^a: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    95
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    96
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    97
lemma (in LP) "A:* P:A->A->* a:A |- Pi a:A. P^a^a: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    98
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
    99
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   100
lemma (in LP) "A:* P:A->* Q:A->* |- Pi a:A. P^a -> Q^a: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   101
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   102
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   103
lemma (in LP) "A:* P:A->* |- Pi a:A. P^a -> P^a: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   104
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   105
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   106
lemma (in LP) "A:* P:A->* |- Lam a:A. Lam x:P^a. x: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   107
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   108
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   109
lemma (in LP) "A:* P:A->* Q:* |- (Pi a:A. P^a->Q) -> (Pi a:A. P^a) -> Q : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   110
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   111
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   112
lemma (in LP) "A:* P:A->* Q:* a0:A |-
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   113
        Lam x:Pi a:A. P^a->Q. Lam y:Pi a:A. P^a. x^a0^(y^a0): ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   114
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   115
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   116
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   117
subsection {* Omega-order types *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   118
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   119
lemma (in L2) "A:* B:* |- Pi C:*.(A->B->C)->C : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   120
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   121
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   122
lemma (in Lomega2) "|- Lam A:*.Lam B:*.Pi C:*.(A->B->C)->C : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   123
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   124
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   125
lemma (in Lomega2) "|- Lam A:*.Lam B:*.Lam x:A. Lam y:B. x : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   126
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   127
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   128
lemma (in Lomega2) "A:* B:* |- ?p : (A->B) -> ((B->Pi P:*.P)->(A->Pi P:*.P))"
17255
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
   129
  apply (strip_asms rules)
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   130
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   131
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   132
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   133
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   134
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   135
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   136
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   137
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   138
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   139
    apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   140
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   141
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   142
  apply (erule pi_elim)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   143
   apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   144
  apply (erule pi_elim)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   145
   apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   146
  apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   147
  done
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   148
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   149
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   150
subsection {* Second-order Predicate Logic *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   151
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   152
lemma (in LP2) "A:* P:A->* |- Lam a:A. P^a->(Pi A:*.A) : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   153
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   154
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   155
lemma (in LP2) "A:* P:A->A->* |-
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   156
    (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   157
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   158
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   159
lemma (in LP2) "A:* P:A->A->* |-
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   160
    ?p: (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   161
  -- {* Antisymmetry implies irreflexivity: *}
17255
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
   162
  apply (strip_asms rules)
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   163
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   164
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   165
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   166
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   167
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   168
    apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   169
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   170
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   171
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   172
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   173
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   174
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   175
  apply (erule pi_elim, assumption, assumption?)+
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   176
  done
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   177
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   178
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   179
subsection {* LPomega *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   180
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   181
lemma (in LPomega) "A:* |- Lam P:A->A->*.Lam a:A. P^a^a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   182
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   183
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   184
lemma (in LPomega) "|- Lam A:*.Lam P:A->A->*.Lam a:A. P^a^a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   185
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   186
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   187
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   188
subsection {* Constructions *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   189
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   190
lemma (in CC) "|- Lam A:*.Lam P:A->*.Lam a:A. P^a->Pi P:*.P: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   191
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   192
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   193
lemma (in CC) "|- Lam A:*.Lam P:A->*.Pi a:A. P^a: ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   194
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   195
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   196
lemma (in CC) "A:* P:A->* a:A |- ?p : (Pi a:A. P^a)->P^a"
17255
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
   197
  apply (strip_asms rules)
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   198
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   199
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   200
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   201
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   202
  apply (erule pi_elim, assumption, assumption)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   203
  done
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   204
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   205
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   206
subsection {* Some random examples *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   207
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   208
lemma (in LP2) "A:* c:A f:A->A |-
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   209
    Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   210
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   211
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   212
lemma (in CC) "Lam A:*.Lam c:A. Lam f:A->A.
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   213
    Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   214
  by (depth_solve rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   215
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   216
lemma (in LP2)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   217
  "A:* a:A b:A |- ?p: (Pi P:A->*.P^a->P^b) -> (Pi P:A->*.P^b->P^a)"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   218
  -- {* Symmetry of Leibnitz equality *}
17255
f9c5cbdce422 tuned method;
wenzelm
parents: 17252
diff changeset
   219
  apply (strip_asms rules)
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   220
  apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   221
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   222
   prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   223
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   224
  apply (erule_tac a = "Lam x:A. Pi Q:A->*.Q^x->Q^a" in pi_elim)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   225
   apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   226
  apply (unfold beta)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   227
  apply (erule imp_elim)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   228
   apply (rule lam_bs)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   229
     apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   230
    prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   231
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   232
   apply (rule lam_ss)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   233
     apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   234
    prefer 2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   235
    apply (depth_solve1 rules)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   236
   apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   237
  apply assumption
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   238
  done
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   239
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 4583
diff changeset
   240
end